3100
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 1
theory TypeSchemes2
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 2
imports "../Nominal2"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 3
begin
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 4
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 5
section {*** Type Schemes defined as a single nominal datatype ***}
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 6
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 7
atom_decl name
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 8
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 9
nominal_datatype ty =
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 10
Var "name"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 11
| Fun "ty" "ty"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 12
and tys =
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 13
All xs::"name fset" ty::"ty" binds (set+) xs in ty
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 14
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 15
thm ty_tys.distinct
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 16
thm ty_tys.induct
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 17
thm ty_tys.inducts
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 18
thm ty_tys.exhaust
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 19
thm ty_tys.strong_exhaust
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 20
thm ty_tys.fv_defs
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 21
thm ty_tys.bn_defs
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 22
thm ty_tys.perm_simps
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 23
thm ty_tys.eq_iff
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 24
thm ty_tys.fv_bn_eqvt
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 25
thm ty_tys.size_eqvt
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 26
thm ty_tys.supports
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 27
thm ty_tys.supp
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 28
thm ty_tys.fresh
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 29
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 30
fun
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 31
lookup :: "(name \<times> ty) list \<Rightarrow> name \<Rightarrow> ty"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 32
where
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 33
"lookup [] Y = Var Y"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 34
| "lookup ((X, T) # Ts) Y = (if X = Y then T else lookup Ts Y)"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 35
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 36
lemma lookup_eqvt[eqvt]:
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 37
shows "(p \<bullet> lookup Ts T) = lookup (p \<bullet> Ts) (p \<bullet> T)"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 38
apply(induct Ts T rule: lookup.induct)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 39
apply(simp_all)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 40
done
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 41
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 42
lemma TEST1:
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 43
assumes "x = Inl y"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 44
shows "(p \<bullet> Sum_Type.Projl x) = Sum_Type.Projl (p \<bullet> x)"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 45
using assms by simp
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 46
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 47
lemma TEST2:
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 48
assumes "x = Inr y"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 49
shows "(p \<bullet> Sum_Type.Projr x) = Sum_Type.Projr (p \<bullet> x)"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 50
using assms by simp
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 51
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 52
lemma test:
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 53
assumes a: "\<exists>y. f x = Inl y"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 54
shows "(p \<bullet> (Sum_Type.Projl (f x))) = Sum_Type.Projl ((p \<bullet> f) (p \<bullet> x))"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 55
using a
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 56
apply clarify
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 57
apply(frule_tac p="p" in permute_boolI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 58
apply(simp (no_asm_use) only: eqvts)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 59
apply(subst (asm) permute_fun_app_eq)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 60
back
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 61
apply(simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 62
done
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 63
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 64
lemma test2:
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 65
assumes a: "\<exists>y. f x = Inl y"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 66
shows "(p \<bullet> (Sum_Type.Projl (f x))) = Sum_Type.Projl (p \<bullet> (f x))"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 67
using a
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 68
apply clarify
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 69
apply(frule_tac p="p" in permute_boolI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 70
apply(simp (no_asm_use) only: eqvts)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 71
apply(subst (asm) permute_fun_app_eq)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 72
back
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 73
apply(simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 74
done
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 75
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 76
nominal_primrec (default "sum_case (\<lambda>x. Inl undefined) (\<lambda>x. Inr undefined)")
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 77
subst :: "(name \<times> ty) list \<Rightarrow> ty \<Rightarrow> ty"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 78
and substs :: "(name \<times> ty) list \<Rightarrow> tys \<Rightarrow> tys"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 79
where
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 80
"subst \<theta> (Var X) = lookup \<theta> X"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 81
| "subst \<theta> (Fun T1 T2) = Fun (subst \<theta> T1) (subst \<theta> T2)"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 82
| "fset (map_fset atom xs) \<sharp>* \<theta> \<Longrightarrow> substs \<theta> (All xs T) = All xs (subst \<theta> T)"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 83
(*unfolding subst_substs_graph_def eqvt_def
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 84
apply rule
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 85
apply perm_simp
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 86
apply (subst test3)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 87
defer
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 88
apply (subst test3)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 89
defer
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 90
apply (subst test3)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 91
defer
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 92
apply rule*)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 93
thm subst_substs_graph.intros
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 94
apply(subgoal_tac "\<And>p x r. subst_substs_graph x r \<Longrightarrow> subst_substs_graph (p \<bullet> x) (p \<bullet> r)")
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 95
apply(simp add: eqvt_def)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 96
apply(rule allI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 97
apply(simp add: permute_fun_def permute_bool_def)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 98
apply(rule ext)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 99
apply(rule ext)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 100
apply(rule iffI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 101
apply(drule_tac x="p" in meta_spec)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 102
apply(drule_tac x="- p \<bullet> x" in meta_spec)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 103
apply(drule_tac x="- p \<bullet> xa" in meta_spec)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 104
apply(simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 105
apply(drule_tac x="-p" in meta_spec)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 106
apply(drule_tac x="x" in meta_spec)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 107
apply(drule_tac x="xa" in meta_spec)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 108
apply(simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 109
--"Eqvt One way"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 110
apply(erule subst_substs_graph.induct)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 111
apply(perm_simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 112
apply(rule subst_substs_graph.intros)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 113
apply(erule subst_substs_graph.cases)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 114
apply(simp (no_asm_use) only: eqvts)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 115
apply(subst test)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 116
back
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 117
apply (rule exI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 118
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 119
apply(rotate_tac 1)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 120
apply(erule subst_substs_graph.cases)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 121
apply(subst test)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 122
back
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 123
apply (rule exI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 124
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 125
apply(perm_simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 126
apply(rule subst_substs_graph.intros)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 127
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 128
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 129
apply(subst test)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 130
back
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 131
apply (rule exI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 132
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 133
apply(perm_simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 134
apply(rule subst_substs_graph.intros)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 135
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 136
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 137
apply(simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 138
--"A"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 139
apply(simp (no_asm_use) only: eqvts)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 140
apply(subst test)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 141
back
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 142
apply (rule exI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 143
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 144
apply(rotate_tac 1)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 145
apply(erule subst_substs_graph.cases)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 146
apply(subst test)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 147
back
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 148
apply (rule exI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 149
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 150
apply(perm_simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 151
apply(rule subst_substs_graph.intros)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 152
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 153
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 154
apply(subst test)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 155
back
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 156
apply (rule exI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 157
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 158
apply(perm_simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 159
apply(rule subst_substs_graph.intros)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 160
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 161
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 162
apply(simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 163
--"A"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 164
apply(simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 165
apply(erule subst_substs_graph.cases)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 166
apply(simp (no_asm_use) only: eqvts)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 167
apply(subst test)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 168
back
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 169
back
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 170
apply (rule exI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 171
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 172
apply(rule subst_substs_graph.intros)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 173
apply (simp add: eqvts)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 174
apply (subgoal_tac "(p \<bullet> (atom ` fset xs)) \<sharp>* (p \<bullet> \<theta>)")
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 175
apply (simp add: image_eqvt eqvts_raw eqvts)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 176
apply (simp add: fresh_star_permute_iff)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 177
apply(perm_simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 178
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 179
apply(simp (no_asm_use) only: eqvts)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 180
apply(subst test)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 181
back
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 182
back
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 183
apply (rule exI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 184
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 185
apply(rule subst_substs_graph.intros)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 186
apply (simp add: eqvts)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 187
apply (subgoal_tac "(p \<bullet> (atom ` fset xs)) \<sharp>* (p \<bullet> \<theta>)")
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 188
apply (simp add: image_eqvt eqvts_raw eqvts)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 189
apply (simp add: fresh_star_permute_iff)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 190
apply(perm_simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 191
apply(assumption)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 192
apply(simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 193
--"Eqvt done"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 194
apply(rule TrueI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 195
apply (case_tac x)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 196
apply simp apply clarify
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 197
apply (rule_tac y="b" in ty_tys.exhaust(1))
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 198
apply (auto)[1]
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 199
apply (auto)[1]
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 200
apply simp apply clarify
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 201
apply (rule_tac ya="b" and c="a" in ty_tys.strong_exhaust(2))
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 202
apply (auto)[1]
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 203
apply (auto)[5]
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 204
--"LAST GOAL"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 205
apply (simp add: meta_eq_to_obj_eq[OF subst_def, symmetric, unfolded fun_eq_iff])
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 206
apply (subgoal_tac "eqvt_at (\<lambda>(l, r). subst l r) (\<theta>', T)")
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 207
apply (thin_tac "eqvt_at subst_substs_sumC (Inl (\<theta>', T))")
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 208
apply (thin_tac "eqvt_at subst_substs_sumC (Inl (\<theta>', Ta))")
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 209
prefer 2
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 210
apply (simp add: eqvt_at_def subst_def)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 211
apply rule
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 212
apply (subst test2)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 213
apply (simp add: subst_substs_sumC_def)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 214
apply (simp add: THE_default_def)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 215
apply (case_tac "Ex1 (subst_substs_graph (Inl (\<theta>', T)))")
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 216
prefer 2
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 217
apply simp
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 218
apply (simp add: the1_equality)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 219
apply auto[1]
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 220
apply (erule_tac x="x" in allE)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 221
apply simp
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 222
apply(cases rule: subst_substs_graph.cases)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 223
apply assumption
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 224
apply (rule_tac x="lookup \<theta> X" in exI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 225
apply clarify
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 226
apply (rule the1_equality)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 227
apply blast apply assumption
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 228
apply (rule_tac x="(Fun (Sum_Type.Projl (subst_substs_sum (Inl (\<theta>, T1))))
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 229
(Sum_Type.Projl (subst_substs_sum (Inl (\<theta>, T2)))))" in exI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 230
apply clarify
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 231
apply (rule the1_equality)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 232
apply blast apply assumption
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 233
apply clarify
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 234
apply simp
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 235
--"-"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 236
apply clarify
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 237
apply (frule supp_eqvt_at)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 238
apply (simp add: finite_supp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 239
apply (erule Abs_res_fcb)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 240
apply (simp add: Abs_fresh_iff)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 241
apply (simp add: Abs_fresh_iff)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 242
apply auto[1]
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 243
apply (simp add: fresh_def fresh_star_def)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 244
apply (erule contra_subsetD)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 245
apply (simp add: supp_Pair)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 246
apply blast
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 247
apply clarify
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 248
apply (simp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 249
apply (simp add: eqvt_at_def)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 250
apply (subst Abs_eq_iff)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 251
apply (rule_tac x="0::perm" in exI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 252
apply (subgoal_tac "p \<bullet> \<theta>' = \<theta>'")
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 253
apply (simp add: alphas fresh_star_zero)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 254
apply (subgoal_tac "\<And>x. x \<in> supp (subst \<theta>' (p \<bullet> T)) \<Longrightarrow> x \<in> p \<bullet> atom ` fset xs \<longleftrightarrow> x \<in> atom ` fset xsa")
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 255
apply blast
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 256
apply (subgoal_tac "x \<in> supp(p \<bullet> \<theta>', p \<bullet> T)")
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 257
apply (simp add: supp_Pair eqvts eqvts_raw)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 258
apply auto[1]
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 259
apply (subgoal_tac "(atom ` fset (p \<bullet> xs)) \<sharp>* \<theta>'")
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 260
apply (simp add: fresh_star_def fresh_def)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 261
apply(drule_tac p1="p" in iffD2[OF fresh_star_permute_iff])
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 262
apply (simp add: eqvts eqvts_raw)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 263
apply (simp add: fresh_star_def fresh_def)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 264
apply (simp (no_asm) only: supp_eqvt[symmetric] Pair_eqvt[symmetric])
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 265
apply (subgoal_tac "p \<bullet> supp (subst \<theta>' T) \<subseteq> p \<bullet> supp (\<theta>', T)")
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 266
apply (erule subsetD)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 267
apply (simp add: supp_eqvt)
3104
f7c4b8e6918b
updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 268
apply (metis permute_pure subset_eqvt)
3100
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 269
apply (rule perm_supp_eq)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 270
apply (simp add: fresh_def fresh_star_def)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 271
apply blast
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 272
done
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 273
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 274
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 275
termination (eqvt) by lexicographic_order
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 276
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 277
text {* Some Tests about Alpha-Equality *}
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 278
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 279
lemma
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 280
shows "All {|a, b|} (Fun (Var a) (Var b)) = All {|b, a|} (Fun (Var a) (Var b))"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 281
apply(simp add: Abs_eq_iff)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 282
apply(rule_tac x="0::perm" in exI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 283
apply(simp add: alphas fresh_star_def ty_tys.supp supp_at_base)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 284
done
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 285
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 286
lemma
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 287
shows "All {|a, b|} (Fun (Var a) (Var b)) = All {|a, b|} (Fun (Var b) (Var a))"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 288
apply(simp add: Abs_eq_iff)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 289
apply(rule_tac x="(atom a \<rightleftharpoons> atom b)" in exI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 290
apply(simp add: alphas fresh_star_def supp_at_base ty_tys.supp)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 291
done
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 292
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 293
lemma
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 294
shows "All {|a, b, c|} (Fun (Var a) (Var b)) = All {|a, b|} (Fun (Var a) (Var b))"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 295
apply(simp add: Abs_eq_iff)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 296
apply(rule_tac x="0::perm" in exI)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 297
apply(simp add: alphas fresh_star_def ty_tys.supp supp_at_base)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 298
done
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 299
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 300
lemma
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 301
assumes a: "a \<noteq> b"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 302
shows "\<not>(All {|a, b|} (Fun (Var a) (Var b)) = All {|c|} (Fun (Var c) (Var c)))"
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 303
using a
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 304
apply(simp add: Abs_eq_iff)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 305
apply(clarify)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 306
apply(simp add: alphas fresh_star_def ty_tys.supp supp_at_base)
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 307
apply auto
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 308
done
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 309
8779fb01d8b4
separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 310
end