Nominal/Ex/TypeSchemes2.thy
author webertj
Wed, 27 Mar 2013 16:09:46 +0100
changeset 3215 3cfd4fc42840
parent 3197 25d11b449e92
child 3229 b52e8651591f
permissions -rw-r--r--
Fixed proofs to work with 13ab4f0a0b0e.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
3215
3cfd4fc42840 Fixed proofs to work with 13ab4f0a0b0e.
webertj
parents: 3197
diff changeset
    56
by (metis Inl_eqvt Projl_Inl permute_fun_app_eq)
3100
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
lemma test2:
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  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
    60
  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
    61
using a
3215
3cfd4fc42840 Fixed proofs to work with 13ab4f0a0b0e.
webertj
parents: 3197
diff changeset
    62
by clarsimp
3100
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
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
    65
    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
    66
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
    67
where
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
  "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
    69
| "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
    70
| "fset (map_fset atom xs) \<sharp>* \<theta> \<Longrightarrow> substs \<theta> (All xs T) = All xs (subst \<theta> T)"
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3183
diff changeset
    71
apply(simp add: subst_substs_graph_aux_def eqvt_def)
3100
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
apply(rule TrueI)
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
apply (case_tac x)
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
apply simp apply clarify 
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
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
    76
apply (auto)[1]
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
apply (auto)[1]
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
apply simp apply clarify 
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
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
    80
apply (auto)[1]
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
apply (auto)[5]
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
--"LAST GOAL"
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
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
    84
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
    85
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
    86
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
    87
prefer 2
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
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
    89
apply rule
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
apply (subst test2)
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
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
    92
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
    93
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
    94
prefer 2
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
apply simp
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
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
    97
apply auto[1]
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
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
    99
apply simp
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
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
   101
apply assumption
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
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
   103
apply clarify
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
apply (rule the1_equality)
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
apply blast apply assumption
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
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
   107
                  (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
   108
apply clarify
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
apply (rule the1_equality)
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
apply blast apply assumption
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
apply clarify
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
apply simp
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
--"-"
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
apply clarify
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
  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
   116
  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
   117
  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
   118
  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
   119
  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
   120
  apply auto[1]
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
  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
   122
  apply (erule contra_subsetD)
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
  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
   124
  apply blast
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
  apply clarify
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
  apply (simp)
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
  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
   128
  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
   129
  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
   130
  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
   131
  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
   132
  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")
3183
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3104
diff changeset
   133
  apply(simp)
3100
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
  apply blast
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
  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
   136
  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
   137
  apply auto[1]
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
  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
   139
  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
   140
  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
   141
  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
   142
  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
   143
  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
   144
  apply (subgoal_tac "p \<bullet> supp (subst \<theta>' T) \<subseteq> p \<bullet> supp (\<theta>', T)")
3183
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3104
diff changeset
   145
  apply (erule_tac subsetD)
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3104
diff changeset
   146
  apply(simp only: supp_eqvt)
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3104
diff changeset
   147
  apply(perm_simp)
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3104
diff changeset
   148
  apply(drule_tac x="p" in spec)
313e6f2cdd89 added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents: 3104
diff changeset
   149
  apply(simp)
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   150
  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
   151
  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
   152
  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
   153
  apply blast
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
  done
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
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
   158
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
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
   160
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
lemma
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
  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
   163
  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
   164
  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
   165
  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
   166
  done
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
lemma
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
  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
   170
  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
   171
  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
   172
  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
   173
  done
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
lemma
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
  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
   177
  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
   178
  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
   179
  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
   180
done
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
lemma
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
  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
   184
  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
   185
  using a
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
  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
   187
  apply(clarify)
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
  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
   189
  apply auto
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
  done
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
end