| author | Christian Urban <urbanc@in.tum.de> | 
| Thu, 24 Feb 2011 16:26:11 +0000 | |
| changeset 2730 | eebc24b9cf39 | 
| parent 2722 | ba34f893539a | 
| child 2727 | c10b56d226ce | 
| permissions | -rw-r--r-- | 
| 1795 | 1 | theory TypeSchemes | 
| 2454 
9ffee4eb1ae1
renamed NewParser to Nominal2
 Christian Urban <urbanc@in.tum.de> parents: 
2451diff
changeset | 2 | imports "../Nominal2" | 
| 1795 | 3 | begin | 
| 4 | ||
| 5 | section {*** Type Schemes ***}
 | |
| 6 | ||
| 2709 | 7 | |
| 2556 
8ed62410236e
added a test about subtyping; disabled two tests, because of problem with function package
 Christian Urban <urbanc@in.tum.de> parents: 
2524diff
changeset | 8 | atom_decl name | 
| 
8ed62410236e
added a test about subtyping; disabled two tests, because of problem with function package
 Christian Urban <urbanc@in.tum.de> parents: 
2524diff
changeset | 9 | |
| 2486 
b4ea19604b0b
cleaned up two examples
 Christian Urban <urbanc@in.tum.de> parents: 
2480diff
changeset | 10 | (* defined as a single nominal datatype *) | 
| 1795 | 11 | |
| 12 | nominal_datatype ty = | |
| 13 | Var "name" | |
| 14 | | Fun "ty" "ty" | |
| 15 | and tys = | |
| 2714 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 16 | All xs::"name fset" ty::"ty" bind (set+) xs in ty | 
| 2434 | 17 | |
| 2468 | 18 | thm ty_tys.distinct | 
| 19 | thm ty_tys.induct | |
| 2617 
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 20 | thm ty_tys.inducts | 
| 
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 21 | thm ty_tys.exhaust ty_tys.strong_exhaust | 
| 2468 | 22 | thm ty_tys.fv_defs | 
| 23 | thm ty_tys.bn_defs | |
| 24 | thm ty_tys.perm_simps | |
| 25 | thm ty_tys.eq_iff | |
| 26 | thm ty_tys.fv_bn_eqvt | |
| 27 | thm ty_tys.size_eqvt | |
| 28 | thm ty_tys.supports | |
| 2493 
2e174807c891
added postprocessed fresh-lemmas for constructors
 Christian Urban <urbanc@in.tum.de> parents: 
2486diff
changeset | 29 | thm ty_tys.supp | 
| 2494 
11133eb76f61
added Foo1 to explore a contrived example
 Christian Urban <urbanc@in.tum.de> parents: 
2493diff
changeset | 30 | thm ty_tys.fresh | 
| 1795 | 31 | |
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 32 | fun | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 33 | lookup :: "(name \<times> ty) list \<Rightarrow> name \<Rightarrow> ty" | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 34 | where | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 35 | "lookup [] Y = Var Y" | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 36 | | "lookup ((X, T) # Ts) Y = (if X = Y then T else lookup Ts Y)" | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 37 | |
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 38 | lemma lookup_eqvt[eqvt]: | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 39 | shows "(p \<bullet> lookup Ts T) = lookup (p \<bullet> Ts) (p \<bullet> T)" | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 40 | apply(induct Ts T rule: lookup.induct) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 41 | apply(simp_all) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 42 | done | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 43 | |
| 2709 | 44 | lemma test: | 
| 45 | assumes a: "f x = Inl y" | |
| 46 | shows "(p \<bullet> (Sum_Type.Projl (f x))) = Sum_Type.Projl ((p \<bullet> f) (p \<bullet> x))" | |
| 47 | using a | |
| 48 | apply(frule_tac p="p" in permute_boolI) | |
| 49 | apply(simp (no_asm_use) only: eqvts) | |
| 50 | apply(subst (asm) permute_fun_app_eq) | |
| 51 | back | |
| 52 | apply(simp) | |
| 53 | done | |
| 54 | ||
| 2710 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 55 | lemma test2: | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 56 | assumes a: "f x = Inl y" | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 57 | shows "(p \<bullet> (Sum_Type.Projl (f x))) = Sum_Type.Projl (p \<bullet> (f x))" | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 58 | using a | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 59 | apply(frule_tac p="p" in permute_boolI) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 60 | apply(simp (no_asm_use) only: eqvts) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 61 | apply(subst (asm) permute_fun_app_eq) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 62 | back | 
| 2709 | 63 | apply(simp) | 
| 2710 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 64 | done | 
| 2709 | 65 | |
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 66 | nominal_primrec | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 67 | subst :: "(name \<times> ty) list \<Rightarrow> ty \<Rightarrow> ty" | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 68 | and substs :: "(name \<times> ty) list \<Rightarrow> tys \<Rightarrow> tys" | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 69 | where | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 70 | "subst \<theta> (Var X) = lookup \<theta> X" | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 71 | | "subst \<theta> (Fun T1 T2) = Fun (subst \<theta> T1) (subst \<theta> T2)" | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 72 | | "fset (map_fset atom xs) \<sharp>* \<theta> \<Longrightarrow> substs \<theta> (All xs T) = All xs (subst \<theta> T)" | 
| 2709 | 73 | |
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 74 | term subst_substs_sumC | 
| 2709 | 75 | thm subst_substs_sumC_def | 
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 76 | term Inl | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 77 | thm subst_substs_graph.induct | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 78 | thm subst_substs_graph.intros | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 79 | thm Projl.simps | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 80 | apply(subgoal_tac "\<And>p x r. subst_substs_graph x r \<Longrightarrow> subst_substs_graph (p \<bullet> x) (p \<bullet> r)") | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 81 | apply(simp add: eqvt_def) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 82 | apply(rule allI) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 83 | apply(simp add: permute_fun_def permute_bool_def) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 84 | apply(rule ext) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 85 | apply(rule ext) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 86 | apply(rule iffI) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 87 | apply(drule_tac x="p" in meta_spec) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 88 | apply(drule_tac x="- p \<bullet> x" in meta_spec) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 89 | apply(drule_tac x="- p \<bullet> xa" in meta_spec) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 90 | apply(simp) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 91 | apply(drule_tac x="-p" in meta_spec) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 92 | apply(drule_tac x="x" in meta_spec) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 93 | apply(drule_tac x="xa" in meta_spec) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 94 | apply(simp) | 
| 2710 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 95 | --"Eqvt One way" | 
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 96 | thm subst_substs_graph.induct | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 97 | thm subst_substs_graph.intros | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 98 | thm Projl.simps | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 99 | apply(erule subst_substs_graph.induct) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 100 | apply(perm_simp) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 101 | apply(rule subst_substs_graph.intros) | 
| 2709 | 102 | thm subst_substs_graph.cases | 
| 103 | apply(erule subst_substs_graph.cases) | |
| 104 | apply(simp (no_asm_use) only: eqvts) | |
| 105 | apply(subst test) | |
| 106 | back | |
| 107 | apply(assumption) | |
| 108 | apply(rotate_tac 1) | |
| 109 | apply(erule subst_substs_graph.cases) | |
| 110 | apply(subst test) | |
| 111 | back | |
| 112 | apply(assumption) | |
| 113 | apply(perm_simp) | |
| 114 | apply(rule subst_substs_graph.intros) | |
| 115 | apply(assumption) | |
| 116 | apply(assumption) | |
| 117 | apply(subst test) | |
| 118 | back | |
| 119 | apply(assumption) | |
| 120 | apply(perm_simp) | |
| 121 | apply(rule subst_substs_graph.intros) | |
| 122 | apply(assumption) | |
| 123 | apply(assumption) | |
| 124 | apply(simp) | |
| 125 | --"A" | |
| 126 | apply(simp (no_asm_use) only: eqvts) | |
| 127 | apply(subst test) | |
| 128 | back | |
| 129 | apply(assumption) | |
| 130 | apply(rotate_tac 1) | |
| 131 | apply(erule subst_substs_graph.cases) | |
| 132 | apply(subst test) | |
| 133 | back | |
| 134 | apply(assumption) | |
| 135 | apply(perm_simp) | |
| 136 | apply(rule subst_substs_graph.intros) | |
| 137 | apply(assumption) | |
| 138 | apply(assumption) | |
| 139 | apply(subst test) | |
| 140 | back | |
| 141 | apply(assumption) | |
| 142 | apply(perm_simp) | |
| 143 | apply(rule subst_substs_graph.intros) | |
| 144 | apply(assumption) | |
| 145 | apply(assumption) | |
| 146 | apply(simp) | |
| 147 | --"A" | |
| 148 | apply(simp) | |
| 149 | apply(erule subst_substs_graph.cases) | |
| 150 | apply(simp (no_asm_use) only: eqvts) | |
| 151 | apply(subst test) | |
| 152 | back | |
| 153 | back | |
| 154 | apply(assumption) | |
| 155 | apply(rule subst_substs_graph.intros) | |
| 2710 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 156 | apply (simp add: eqvts) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 157 | apply (subgoal_tac "(p \<bullet> (atom ` fset xs)) \<sharp>* (p \<bullet> \<theta>)") | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 158 | apply (simp add: image_eqvt eqvts_raw eqvts) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 159 | apply (simp add: fresh_star_permute_iff) | 
| 2709 | 160 | apply(perm_simp) | 
| 161 | apply(assumption) | |
| 162 | apply(simp (no_asm_use) only: eqvts) | |
| 163 | apply(subst test) | |
| 164 | back | |
| 165 | back | |
| 166 | apply(assumption) | |
| 167 | apply(rule subst_substs_graph.intros) | |
| 2710 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 168 | apply (simp add: eqvts) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 169 | apply (subgoal_tac "(p \<bullet> (atom ` fset xs)) \<sharp>* (p \<bullet> \<theta>)") | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 170 | apply (simp add: image_eqvt eqvts_raw eqvts) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 171 | apply (simp add: fresh_star_permute_iff) | 
| 2709 | 172 | apply(perm_simp) | 
| 173 | apply(assumption) | |
| 174 | apply(simp) | |
| 2710 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 175 | --"Eqvt done" | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 176 | apply (case_tac x) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 177 | apply simp apply clarify | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 178 | apply (rule_tac y="b" in ty_tys.exhaust(1)) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 179 | apply (auto simp add: ty_tys.eq_iff)[1] | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 180 | apply (auto simp add: ty_tys.eq_iff)[1] | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 181 | apply blast | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 182 | apply simp apply clarify | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 183 | apply (rule_tac ya="b" and c="a" in ty_tys.strong_exhaust(2)) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 184 | apply (auto simp add: ty_tys.eq_iff)[1] | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 185 | apply (auto simp add: ty_tys.distinct) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 186 | apply (auto simp add: ty_tys.eq_iff ty_tys.distinct)[2] | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 187 | --"LAST GOAL" | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 188 | thm meta_eq_to_obj_eq[OF subst_def, symmetric, unfolded fun_eq_iff] | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 189 | apply (simp add: meta_eq_to_obj_eq[OF subst_def, symmetric, unfolded fun_eq_iff]) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 190 | apply (subgoal_tac "eqvt_at (\<lambda>(l, r). subst l r) (\<theta>', T)") | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 191 | apply (thin_tac "eqvt_at subst_substs_sumC (Inl (\<theta>', T))") | 
| 2709 | 192 | defer | 
| 2710 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 193 | apply (simp add: eqvt_at_def subst_def) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 194 | apply rule | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 195 | apply (subgoal_tac "\<And>x. subst_substs_sumC (Inl (x)) = Inl (?y x)") | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 196 | apply (subst test2) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 197 | apply (drule_tac x="(\<theta>', T)" in meta_spec) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 198 | apply assumption | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 199 | apply simp | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 200 | --"We require that for Inl it returns Inl. It doesn't work for undefined, but it does work for the following" | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 201 | apply (subgoal_tac "\<And>y. \<exists>z. (\<lambda>x. THE_default (sum_case (\<lambda>x. Inl undefined) (\<lambda>x. Inr undefined) x) (subst_substs_graph x)) (Inl y) = (Inl z)") | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 202 | prefer 2 | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 203 | apply (simp add: THE_default_def) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 204 | apply (case_tac "Ex1 (subst_substs_graph (Inl y))") | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 205 | prefer 2 | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 206 | apply simp | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 207 | apply (simp add: the1_equality) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 208 | apply auto[1] | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 209 | apply (erule_tac x="x" in allE) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 210 | apply simp | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 211 | apply(cases rule: subst_substs_graph.cases) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 212 | apply assumption | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 213 | apply (rule_tac x="lookup \<theta> X" in exI) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 214 | apply clarify | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 215 | apply (rule the1_equality) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 216 | apply metis apply assumption | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 217 | apply (rule_tac x="(Fun (Sum_Type.Projl (subst_substs_sum (Inl (\<theta>, T1)))) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 218 | (Sum_Type.Projl (subst_substs_sum (Inl (\<theta>, T2)))))" in exI) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 219 | apply clarify | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 220 | apply (rule the1_equality) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 221 | apply metis apply assumption | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 222 | apply clarify | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 223 | --"This is exactly the assumption for the properly defined function" | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 224 | defer | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 225 | apply (simp add: ty_tys.eq_iff) | 
| 2714 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 226 | apply (simp only: Abs_eq_res_set) | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 227 | apply (subgoal_tac "(atom ` fset xsa \<inter> supp T - atom ` fset xs \<inter> supp Ta) \<sharp>* ([atom ` fset xs \<inter> supp (subst \<theta>' T)]set. T)") | 
| 2710 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 228 | apply (subst (asm) Abs_eq_iff2) | 
| 2714 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 229 | apply (clarify) | 
| 2711 
ec1a7ef740b8
Showing that the binders difference is fresh for the left side solves the goal for 'set'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2710diff
changeset | 230 | apply (simp add: alphas) | 
| 2710 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 231 | apply (clarify) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 232 | apply (rule trans) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 233 | apply(rule_tac p="p" in supp_perm_eq[symmetric]) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 234 | apply(rule fresh_star_supp_conv) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 235 | thm fresh_star_perm_set_conv | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 236 | apply(drule fresh_star_perm_set_conv) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 237 | apply (rule finite_Diff) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 238 | apply (rule finite_supp) | 
| 2714 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 239 | apply (subgoal_tac "(atom ` fset xs \<inter> supp T \<union> atom ` fset xsa \<inter> supp (p \<bullet> T)) \<sharp>* ([atom ` fset xs \<inter> supp (subst \<theta>' T)]set. subst \<theta>' T)") | 
| 2710 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 240 | apply (metis Un_absorb2 fresh_star_Un) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 241 | apply (simp add: fresh_star_Un) | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 242 | apply (rule conjI) | 
| 2714 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 243 | apply (simp (no_asm) add: fresh_star_def) | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 244 | |
| 2711 
ec1a7ef740b8
Showing that the binders difference is fresh for the left side solves the goal for 'set'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2710diff
changeset | 245 | apply rule | 
| 
ec1a7ef740b8
Showing that the binders difference is fresh for the left side solves the goal for 'set'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2710diff
changeset | 246 | apply(simp (no_asm) only: Abs_fresh_iff) | 
| 
ec1a7ef740b8
Showing that the binders difference is fresh for the left side solves the goal for 'set'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2710diff
changeset | 247 | apply(clarify) | 
| 2714 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 248 | apply auto[1] | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 249 | apply (simp add: fresh_star_def fresh_def) | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 250 | --"HERE" | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 251 | apply (simp (no_asm) add: fresh_star_def) | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 252 | apply rule | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 253 | apply auto[1] | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 254 | apply(simp (no_asm) only: Abs_fresh_iff) | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 255 | apply(clarify) | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 256 | apply auto[1] | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 257 | prefer 2 | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 258 | apply (simp add: fresh_def) | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 259 | apply(drule_tac a="atom x" in fresh_eqvt_at) | 
| 2711 
ec1a7ef740b8
Showing that the binders difference is fresh for the left side solves the goal for 'set'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2710diff
changeset | 260 | apply (simp add: supp_Pair finite_supp) | 
| 
ec1a7ef740b8
Showing that the binders difference is fresh for the left side solves the goal for 'set'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2710diff
changeset | 261 | apply (simp add: fresh_Pair) | 
| 2714 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 262 | apply(auto simp add: Abs_fresh_iff fresh_star_def)[1] | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 263 | prefer 2 | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 264 | apply auto[1] | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 265 | apply (erule_tac x="atom x" in ballE) | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 266 | apply auto[1] | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 267 | apply (auto simp add: fresh_def)[1] | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 268 | |
| 2710 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 269 | apply (subgoal_tac "p \<bullet> \<theta>' = \<theta>'") | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 270 | prefer 2 | 
| 
7eebe0d5d298
Experiments with functions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2709diff
changeset | 271 | apply (rule perm_supp_eq) | 
| 2714 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 272 | apply (subgoal_tac "(atom ` fset xs \<inter> supp T \<union> atom ` fset xsa \<inter> supp (p \<bullet> T)) \<sharp>* \<theta>'") | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 273 | apply (auto simp add: fresh_star_def)[1] | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 274 | apply (simp add: fresh_star_Un fresh_star_def) | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 275 | apply blast | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 276 | apply(simp add: eqvt_at_def inter_eqvt supp_eqvt) | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 277 | apply (simp only: Abs_eq_res_set[symmetric]) | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 278 | |
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 279 | apply (rule_tac s="[p \<bullet> atom ` fset xs \<inter> supp (\<theta>', p \<bullet> T)]res. subst \<theta>' (p \<bullet> T)" in trans) | 
| 
908750991c2f
Experiments with substitution on set+
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2711diff
changeset | 280 | --"What if (p \<bullet> xs) is not fresh for \<theta>' ?" | 
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 281 | oops | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 282 | |
| 2676 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 283 | |
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 284 | section {* defined as two separate nominal datatypes *}
 | 
| 2486 
b4ea19604b0b
cleaned up two examples
 Christian Urban <urbanc@in.tum.de> parents: 
2480diff
changeset | 285 | |
| 2308 
387fcbd33820
fixed problem with bn_info
 Christian Urban <urbanc@in.tum.de> parents: 
2181diff
changeset | 286 | nominal_datatype ty2 = | 
| 
387fcbd33820
fixed problem with bn_info
 Christian Urban <urbanc@in.tum.de> parents: 
2181diff
changeset | 287 | Var2 "name" | 
| 
387fcbd33820
fixed problem with bn_info
 Christian Urban <urbanc@in.tum.de> parents: 
2181diff
changeset | 288 | | Fun2 "ty2" "ty2" | 
| 
387fcbd33820
fixed problem with bn_info
 Christian Urban <urbanc@in.tum.de> parents: 
2181diff
changeset | 289 | |
| 
387fcbd33820
fixed problem with bn_info
 Christian Urban <urbanc@in.tum.de> parents: 
2181diff
changeset | 290 | nominal_datatype tys2 = | 
| 2634 
3ce1089cdbf3
changed res keyword to set+ for restrictions; comment by a referee
 Christian Urban <urbanc@in.tum.de> parents: 
2630diff
changeset | 291 | All2 xs::"name fset" ty::"ty2" bind (set+) xs in ty | 
| 2337 
b151399bd2c3
fixed according to changes in quotient
 Christian Urban <urbanc@in.tum.de> parents: 
2308diff
changeset | 292 | |
| 2468 | 293 | thm tys2.distinct | 
| 2630 
8268b277d240
automated all strong induction lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2622diff
changeset | 294 | thm tys2.induct tys2.strong_induct | 
| 2617 
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 295 | thm tys2.exhaust tys2.strong_exhaust | 
| 2468 | 296 | thm tys2.fv_defs | 
| 297 | thm tys2.bn_defs | |
| 298 | thm tys2.perm_simps | |
| 299 | thm tys2.eq_iff | |
| 300 | thm tys2.fv_bn_eqvt | |
| 301 | thm tys2.size_eqvt | |
| 302 | thm tys2.supports | |
| 2493 
2e174807c891
added postprocessed fresh-lemmas for constructors
 Christian Urban <urbanc@in.tum.de> parents: 
2486diff
changeset | 303 | thm tys2.supp | 
| 2494 
11133eb76f61
added Foo1 to explore a contrived example
 Christian Urban <urbanc@in.tum.de> parents: 
2493diff
changeset | 304 | thm tys2.fresh | 
| 2468 | 305 | |
| 2676 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 306 | fun | 
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 307 | lookup2 :: "(name \<times> ty2) list \<Rightarrow> name \<Rightarrow> ty2" | 
| 2676 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 308 | where | 
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 309 | "lookup2 [] Y = Var2 Y" | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 310 | | "lookup2 ((X, T) # Ts) Y = (if X = Y then T else lookup2 Ts Y)" | 
| 2556 
8ed62410236e
added a test about subtyping; disabled two tests, because of problem with function package
 Christian Urban <urbanc@in.tum.de> parents: 
2524diff
changeset | 311 | |
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 312 | lemma lookup2_eqvt[eqvt]: | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 313 | shows "(p \<bullet> lookup2 Ts T) = lookup2 (p \<bullet> Ts) (p \<bullet> T)" | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 314 | apply(induct Ts T rule: lookup2.induct) | 
| 2676 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 315 | apply(simp_all) | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 316 | done | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 317 | |
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 318 | nominal_primrec | 
| 2676 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 319 | subst :: "(name \<times> ty2) list \<Rightarrow> ty2 \<Rightarrow> ty2" | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 320 | where | 
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 321 | "subst \<theta> (Var2 X) = lookup2 \<theta> X" | 
| 2676 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 322 | | "subst \<theta> (Fun2 T1 T2) = Fun2 (subst \<theta> T1) (subst \<theta> T2)" | 
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 323 | defer | 
| 2676 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 324 | apply(case_tac x) | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 325 | apply(simp) | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 326 | apply(rule_tac y="b" in ty2.exhaust) | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 327 | apply(blast) | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 328 | apply(blast) | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 329 | apply(simp_all add: ty2.distinct) | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 330 | apply(simp add: ty2.eq_iff) | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 331 | apply(simp add: ty2.eq_iff) | 
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 332 | apply(subgoal_tac "\<And>p x r. subst_graph x r \<Longrightarrow> subst_graph (p \<bullet> x) (p \<bullet> r)") | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 333 | apply(simp add: eqvt_def) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 334 | apply(rule allI) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 335 | apply(simp add: permute_fun_def permute_bool_def) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 336 | apply(rule ext) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 337 | apply(rule ext) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 338 | apply(rule iffI) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 339 | apply(drule_tac x="p" in meta_spec) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 340 | apply(drule_tac x="- p \<bullet> x" in meta_spec) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 341 | apply(drule_tac x="- p \<bullet> xa" in meta_spec) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 342 | apply(simp) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 343 | apply(drule_tac x="-p" in meta_spec) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 344 | apply(drule_tac x="x" in meta_spec) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 345 | apply(drule_tac x="xa" in meta_spec) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 346 | apply(simp) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 347 | apply(erule subst_graph.induct) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 348 | apply(perm_simp) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 349 | apply(rule subst_graph.intros) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 350 | apply(perm_simp) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 351 | apply(rule subst_graph.intros) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 352 | apply(assumption) | 
| 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 353 | apply(assumption) | 
| 2676 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 354 | done | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 355 | |
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 356 | termination | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 357 | apply(relation "measure (size o snd)") | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 358 | apply(simp_all add: ty2.size) | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 359 | done | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 360 | |
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 361 | lemma subst_eqvt[eqvt]: | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 362 | shows "(p \<bullet> subst \<theta> T) = subst (p \<bullet> \<theta>) (p \<bullet> T)" | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 363 | apply(induct \<theta> T rule: subst.induct) | 
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 364 | apply(simp_all add: lookup2_eqvt) | 
| 2676 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 365 | done | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 366 | |
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 367 | lemma j: | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 368 | assumes "a \<sharp> Ts" " a \<sharp> X" | 
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 369 | shows "a \<sharp> lookup2 Ts X" | 
| 2676 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 370 | using assms | 
| 2707 
747ebf2f066d
made eqvt-proof explicit in the function definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2676diff
changeset | 371 | apply(induct Ts X rule: lookup2.induct) | 
| 2676 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 372 | apply(auto simp add: ty2.fresh fresh_Cons fresh_Pair) | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 373 | done | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 374 | |
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 375 | lemma i: | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 376 | assumes "a \<sharp> t" " a \<sharp> \<theta>" | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 377 | shows "a \<sharp> subst \<theta> t" | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 378 | using assms | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 379 | apply(induct \<theta> t rule: subst.induct) | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 380 | apply(auto simp add: ty2.fresh j) | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 381 | done | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 382 | |
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 383 | lemma k: | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 384 | assumes "as \<sharp>* t" " as \<sharp>* \<theta>" | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 385 | shows "as \<sharp>* subst \<theta> t" | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 386 | using assms | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 387 | by (simp add: fresh_star_def i) | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 388 | |
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 389 | lemma h: | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 390 | assumes "as \<subseteq> bs \<union> cs" | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 391 | and " cs \<sharp>* x" | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 392 | shows "(as - bs) \<sharp>* x" | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 393 | using assms | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 394 | by (auto simp add: fresh_star_def) | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 395 | |
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 396 | nominal_primrec | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 397 | substs :: "(name \<times> ty2) list \<Rightarrow> tys2 \<Rightarrow> tys2" | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 398 | where | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 399 | "fset (map_fset atom xs) \<sharp>* \<theta> \<Longrightarrow> substs \<theta> (All2 xs t) = All2 xs (subst \<theta> t)" | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 400 | oops | 
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 401 | |
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 402 | |
| 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 403 | text {* Some Tests about Alpha-Equality *}
 | 
| 1795 | 404 | |
| 405 | lemma | |
| 406 |   shows "All {|a, b|} (Fun (Var a) (Var b)) = All {|b, a|} (Fun (Var a) (Var b))"
 | |
| 2676 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 407 | apply(simp add: ty_tys.eq_iff Abs_eq_iff) | 
| 1795 | 408 | apply(rule_tac x="0::perm" in exI) | 
| 2676 
028d5511c15f
some tryes about substitution over type-schemes
 Christian Urban <urbanc@in.tum.de> parents: 
2634diff
changeset | 409 | apply(simp add: alphas fresh_star_def ty_tys.supp supp_at_base) | 
| 1795 | 410 | done | 
| 411 | ||
| 412 | lemma | |
| 413 |   shows "All {|a, b|} (Fun (Var a) (Var b)) = All {|a, b|} (Fun (Var b) (Var a))"
 | |
| 2566 
a59d8e1e3a17
moved rest of the lemmas from Nominal2_FSet to the TypeScheme example
 Christian Urban <urbanc@in.tum.de> parents: 
2556diff
changeset | 414 | apply(simp add: ty_tys.eq_iff Abs_eq_iff) | 
| 1795 | 415 | apply(rule_tac x="(atom a \<rightleftharpoons> atom b)" in exI) | 
| 2566 
a59d8e1e3a17
moved rest of the lemmas from Nominal2_FSet to the TypeScheme example
 Christian Urban <urbanc@in.tum.de> parents: 
2556diff
changeset | 416 | apply(simp add: alphas fresh_star_def supp_at_base ty_tys.supp) | 
| 1795 | 417 | done | 
| 418 | ||
| 419 | lemma | |
| 420 |   shows "All {|a, b, c|} (Fun (Var a) (Var b)) = All {|a, b|} (Fun (Var a) (Var b))"
 | |
| 2566 
a59d8e1e3a17
moved rest of the lemmas from Nominal2_FSet to the TypeScheme example
 Christian Urban <urbanc@in.tum.de> parents: 
2556diff
changeset | 421 | apply(simp add: ty_tys.eq_iff Abs_eq_iff) | 
| 1795 | 422 | apply(rule_tac x="0::perm" in exI) | 
| 2566 
a59d8e1e3a17
moved rest of the lemmas from Nominal2_FSet to the TypeScheme example
 Christian Urban <urbanc@in.tum.de> parents: 
2556diff
changeset | 423 | apply(simp add: alphas fresh_star_def ty_tys.supp supp_at_base) | 
| 1795 | 424 | done | 
| 425 | ||
| 426 | lemma | |
| 427 | assumes a: "a \<noteq> b" | |
| 428 |   shows "\<not>(All {|a, b|} (Fun (Var a) (Var b)) = All {|c|} (Fun (Var c) (Var c)))"
 | |
| 429 | using a | |
| 2566 
a59d8e1e3a17
moved rest of the lemmas from Nominal2_FSet to the TypeScheme example
 Christian Urban <urbanc@in.tum.de> parents: 
2556diff
changeset | 430 | apply(simp add: ty_tys.eq_iff Abs_eq_iff) | 
| 1795 | 431 | apply(clarify) | 
| 2566 
a59d8e1e3a17
moved rest of the lemmas from Nominal2_FSet to the TypeScheme example
 Christian Urban <urbanc@in.tum.de> parents: 
2556diff
changeset | 432 | apply(simp add: alphas fresh_star_def ty_tys.eq_iff ty_tys.supp supp_at_base) | 
| 1795 | 433 | apply auto | 
| 434 | done | |
| 435 | ||
| 2566 
a59d8e1e3a17
moved rest of the lemmas from Nominal2_FSet to the TypeScheme example
 Christian Urban <urbanc@in.tum.de> parents: 
2556diff
changeset | 436 | |
| 1795 | 437 | |
| 438 | ||
| 439 | end |