Nominal/Ex/TypeSchemes1.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 04 Jan 2012 17:42:16 +0000
changeset 3105 1b0d230445ce
parent 3104 f7c4b8e6918b
child 3109 d79e936e30ea
permissions -rw-r--r--
added an FCB for res (will not define evry function, but is a good datapoint)
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: 2982
diff changeset
     1
theory TypeSchemes1
2454
9ffee4eb1ae1 renamed NewParser to Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 2451
diff changeset
     2
imports "../Nominal2"
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
     5
section {* Type Schemes defined as two separate nominal datatypes *}
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
2556
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2524
diff changeset
     7
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: 2524
diff changeset
     8
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
nominal_datatype ty =
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
  Var "name"
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
    11
| Fun "ty" "ty" ("_ \<rightarrow> _")
3100
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    12
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    13
nominal_datatype tys =
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
    14
  All xs::"name fset" ty::"ty" binds (set+) xs in ty ("All [_]._")
2434
92dc6cfa3a95 automatic lifting
Christian Urban <urbanc@in.tum.de>
parents: 2424
diff changeset
    15
3100
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    16
thm tys.distinct
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    17
thm tys.induct tys.strong_induct
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    18
thm tys.exhaust tys.strong_exhaust
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    19
thm tys.fv_defs
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    20
thm tys.bn_defs
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    21
thm tys.perm_simps
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    22
thm tys.eq_iff
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    23
thm tys.fv_bn_eqvt
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    24
thm tys.size_eqvt
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    25
thm tys.supports
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    26
thm tys.supp
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    27
thm tys.fresh
2885
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    28
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    29
subsection {* Some Tests about Alpha-Equality *}
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    30
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    31
lemma
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    32
  shows "All [{|a, b|}].((Var a) \<rightarrow> (Var b)) = All [{|b, a|}]. ((Var a) \<rightarrow> (Var b))"
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    33
  apply(simp add: Abs_eq_iff)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    34
  apply(rule_tac x="0::perm" in exI)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    35
  apply(simp add: alphas fresh_star_def ty.supp supp_at_base)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    36
  done
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    37
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    38
lemma
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    39
  shows "All [{|a, b|}].((Var a) \<rightarrow> (Var b)) = All [{|a, b|}].((Var b) \<rightarrow> (Var a))"
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    40
  apply(simp add: Abs_eq_iff)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    41
  apply(rule_tac x="(atom a \<rightleftharpoons> atom b)" in exI)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    42
  apply(simp add: alphas fresh_star_def supp_at_base ty.supp)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    43
  done
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    44
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    45
lemma
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    46
  shows "All [{|a, b, c|}].((Var a) \<rightarrow> (Var b)) = All [{|a, b|}].((Var a) \<rightarrow> (Var b))"
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    47
  apply(simp add: Abs_eq_iff)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    48
  apply(rule_tac x="0::perm" in exI)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    49
  apply(simp add: alphas fresh_star_def ty.supp supp_at_base)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    50
done
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    51
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    52
lemma
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    53
  assumes a: "a \<noteq> b"
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    54
  shows "\<not>(All [{|a, b|}].((Var a) \<rightarrow> (Var b)) = All [{|c|}].((Var c) \<rightarrow> (Var c)))"
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    55
  using a
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    56
  apply(simp add: Abs_eq_iff)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    57
  apply(clarify)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    58
  apply(simp add: alphas fresh_star_def ty.supp supp_at_base)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    59
  apply auto
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    60
  done
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    61
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    62
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
    63
subsection {* Substitution function for types and type schemes *}
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
    64
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
    65
type_synonym 
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
    66
  Subst = "(name \<times> ty) list"
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    68
fun
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
    69
  lookup :: "Subst \<Rightarrow> name \<Rightarrow> ty"
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    70
where
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    71
  "lookup [] Y = Var Y"
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    72
| "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: 2676
diff changeset
    73
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    74
lemma lookup_eqvt[eqvt]:
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    75
  shows "(p \<bullet> lookup Ts T) = lookup (p \<bullet> Ts) (p \<bullet> T)"
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
    76
  by (induct Ts T rule: lookup.induct) (simp_all)
2981
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
    77
3100
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    78
nominal_primrec
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
    79
  subst  :: "Subst \<Rightarrow> ty \<Rightarrow> ty" ("_<_>" [100,60] 120)
2838
36544bac1659 More experiments with 'default'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2836
diff changeset
    80
where
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
    81
  "\<theta><Var X> = lookup \<theta> X"
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
    82
| "\<theta><T1 \<rightarrow> T2> = (\<theta><T1>) \<rightarrow> (\<theta><T2>)"
3100
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    83
  unfolding eqvt_def subst_graph_def
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
    84
  apply (rule, perm_simp, rule)
2822
23befefc6e73 cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
Christian Urban <urbanc@in.tum.de>
parents: 2805
diff changeset
    85
  apply(rule TrueI)
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
    86
  apply(case_tac x)
3100
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    87
  apply(rule_tac y="b" in ty.exhaust)
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
    88
  apply(blast)
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
    89
  apply(blast)
3100
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    90
  apply(simp_all)
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
    91
  done
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
    92
2973
d1038e67923a added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
parents: 2950
diff changeset
    93
termination (eqvt)
2859
2eeb75cccb8d all tests work again
Christian Urban <urbanc@in.tum.de>
parents: 2850
diff changeset
    94
  by lexicographic_order
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
    95
3100
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
    96
lemma supp_fun_app_eqvt:
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
    97
  assumes e: "eqvt f"
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
    98
  shows "supp (f a b) \<subseteq> supp a \<union> supp b"
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
    99
  using supp_fun_app_eqvt[OF e] supp_fun_app
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   100
  by blast
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   101
 
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   102
lemma supp_subst:
3100
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
   103
  "supp (subst \<theta> t) \<subseteq> supp \<theta> \<union> supp t"
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
   104
  apply (rule supp_fun_app_eqvt)
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
   105
  unfolding eqvt_def 
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
   106
  by (simp add: permute_fun_def subst.eqvt)
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   107
 
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   108
nominal_primrec
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   109
  substs :: "(name \<times> ty) list \<Rightarrow> tys \<Rightarrow> tys" ("_<_>" [100,60] 120)
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   110
where
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   111
  "fset (map_fset atom Xs) \<sharp>* \<theta> \<Longrightarrow> \<theta><All [Xs].T> = All [Xs].(\<theta><T>)"
3100
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
   112
  unfolding eqvt_def substs_graph_def
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   113
  apply (rule, perm_simp, rule)
2822
23befefc6e73 cleaned ups a bit the examples with the invariant framework; exported nominal_function_config datatype into separate structure and file
Christian Urban <urbanc@in.tum.de>
parents: 2805
diff changeset
   114
  apply auto[2]
3100
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
   115
  apply (rule_tac y="b" and c="a" in tys.strong_exhaust)
2982
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   116
  apply auto[1]
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   117
  apply(simp)
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   118
  apply(erule conjE)
2830
297cff1d1048 FCB for res binding and simplified proof of subst for type schemes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2822
diff changeset
   119
  apply (erule Abs_res_fcb)
297cff1d1048 FCB for res binding and simplified proof of subst for type schemes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2822
diff changeset
   120
  apply (simp add: Abs_fresh_iff)
2982
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   121
  apply(simp add: fresh_def)
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   122
  apply(simp add: supp_Abs)
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   123
  apply(rule impI)
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   124
  apply(subgoal_tac "x \<notin> supp \<theta>")
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   125
  prefer 2
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   126
  apply(auto simp add: fresh_star_def fresh_def)[1]
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   127
  apply(subgoal_tac "x \<in> supp T")
2982
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   128
  using supp_subst
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   129
  apply(blast)
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   130
  using supp_subst
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   131
  apply(blast)
2832
76db0b854bf6 Simpler proof of TypeSchemes/substs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2831
diff changeset
   132
  apply clarify
3100
8779fb01d8b4 separated the two versions of type schemes into two files
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
   133
  apply (simp add: subst.eqvt)
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   134
  apply (subst Abs_eq_iff)
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   135
  apply (rule_tac x="0::perm" in exI)
2832
76db0b854bf6 Simpler proof of TypeSchemes/substs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2831
diff changeset
   136
  apply (subgoal_tac "p \<bullet> \<theta>' = \<theta>'")
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   137
  apply (simp add: alphas fresh_star_zero)
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   138
  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")
2804
db0ed02eba6e Remove SMT
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2801
diff changeset
   139
  apply blast
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   140
  apply (subgoal_tac "x \<in> supp(p \<bullet> \<theta>', p \<bullet> T)")
2839
bcf48a1cb24b abs_res_fcb will be enough to finish the multiple-recursive proof, if we have a working 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2838
diff changeset
   141
  apply (simp add: supp_Pair eqvts eqvts_raw)
bcf48a1cb24b abs_res_fcb will be enough to finish the multiple-recursive proof, if we have a working 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2838
diff changeset
   142
  apply auto[1]
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   143
  apply (subgoal_tac "(atom ` fset (p \<bullet> Xs)) \<sharp>* \<theta>'")
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   144
  apply (simp add: fresh_star_def fresh_def)
2839
bcf48a1cb24b abs_res_fcb will be enough to finish the multiple-recursive proof, if we have a working 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2838
diff changeset
   145
  apply(drule_tac p1="p" in iffD2[OF fresh_star_permute_iff])
bcf48a1cb24b abs_res_fcb will be enough to finish the multiple-recursive proof, if we have a working 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2838
diff changeset
   146
  apply (simp add: eqvts eqvts_raw)
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   147
  apply (simp add: fresh_star_def fresh_def)
2839
bcf48a1cb24b abs_res_fcb will be enough to finish the multiple-recursive proof, if we have a working 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2838
diff changeset
   148
  apply (drule subsetD[OF supp_subst])
bcf48a1cb24b abs_res_fcb will be enough to finish the multiple-recursive proof, if we have a working 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2838
diff changeset
   149
  apply (simp add: supp_Pair)
2832
76db0b854bf6 Simpler proof of TypeSchemes/substs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2831
diff changeset
   150
  apply (rule perm_supp_eq)
76db0b854bf6 Simpler proof of TypeSchemes/substs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2831
diff changeset
   151
  apply (simp add: fresh_def fresh_star_def)
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   152
  apply blast
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   153
  done
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   154
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   155
termination (eqvt)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   156
  by (lexicographic_order)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   157
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   159
subsection {* Generalisation of types and type-schemes*}
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   161
fun
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   162
  subst_dom_pi :: "Subst \<Rightarrow> perm \<Rightarrow> Subst" ("_|_")
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   163
where
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   164
  "[]|p = []"
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   165
| "((X,T)#\<theta>)|p = (p \<bullet> X, T)#(\<theta>|p)" 
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   166
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   167
fun
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   168
  subst_subst :: "Subst \<Rightarrow> Subst \<Rightarrow> Subst" ("_<_>"  [100,60] 120)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   169
where
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   170
  "\<theta><[]> = []"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   171
| "\<theta> <((X,T)#\<theta>')> = (X,\<theta><T>)#(\<theta><\<theta>'>)"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   172
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   173
fun
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   174
  dom :: "Subst \<Rightarrow> name fset"
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   175
where
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   176
  "dom [] = {||}"
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   177
| "dom ((X,T)#\<theta>) = {|X|} |\<union>| dom \<theta>"
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   178
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   179
lemma dom_eqvt[eqvt]:
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   180
  shows "p \<bullet> (dom \<theta>) = dom (p \<bullet> \<theta>)"
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   181
by (induct \<theta>) (auto)
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   182
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   183
lemma dom_subst:
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   184
  fixes \<theta>1 \<theta>2::"Subst"
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   185
  shows "dom (\<theta>2<\<theta>1>) = (dom \<theta>1)"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   186
by (induct \<theta>1) (auto)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   187
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   188
lemma dom_pi:
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   189
  shows "dom (\<theta>|p) = dom (p \<bullet> \<theta>)"
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   190
by (induct \<theta>) (auto)
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   191
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   192
lemma dom_fresh_lookup:
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   193
  fixes \<theta>::"Subst"
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   194
  assumes "\<forall>Y \<in> fset (dom \<theta>). atom Y \<sharp> X"
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   195
  shows "lookup \<theta> X = Var X"
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   196
using assms
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   197
by (induct \<theta>) (auto simp add: fresh_at_base)
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   198
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   199
lemma dom_fresh_ty:
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   200
  fixes \<theta>::"Subst"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   201
  and   T::"ty"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   202
  assumes "\<forall>X \<in> fset (dom \<theta>). atom X \<sharp> T"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   203
  shows "\<theta><T> = T"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   204
using assms
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   205
by (induct T rule: ty.induct) (auto simp add: ty.fresh dom_fresh_lookup)
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   206
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   207
lemma dom_fresh_subst:
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   208
  fixes \<theta> \<theta>'::"Subst"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   209
  assumes "\<forall>X \<in> fset (dom \<theta>). atom X \<sharp> \<theta>'"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   210
  shows "\<theta><\<theta>'> = \<theta>'"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   211
using assms
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   212
by (induct \<theta>') (auto simp add: fresh_Pair fresh_Cons dom_fresh_ty)
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   213
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   214
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   215
definition
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   216
  generalises :: "ty \<Rightarrow> tys \<Rightarrow> bool" ("_ \<prec>\<prec> _")
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   217
where
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   218
  "T \<prec>\<prec> S \<longleftrightarrow> (\<exists>\<theta> Xs T'. S = All [Xs].T'\<and> T = \<theta><T'> \<and> dom \<theta> = Xs)"
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   219
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   220
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   221
lemma lookup_fresh:
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   222
  fixes X::"name"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   223
  assumes a: "atom X \<sharp> \<theta>"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   224
  shows "lookup \<theta> X = Var X"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   225
  using a
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   226
  by (induct \<theta>) (auto simp add: fresh_Cons fresh_Pair fresh_at_base)
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   227
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   228
lemma lookup_dom:
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   229
  fixes X::"name"
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   230
  assumes "X |\<notin>| dom \<theta>"
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   231
  shows "lookup \<theta> X = Var X"
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   232
  using assms
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   233
  by (induct \<theta>) (auto)
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   234
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   235
lemma w1: 
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   236
  shows "\<theta><\<theta>'|p> = (\<theta><\<theta>'>)|p"
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   237
  by (induct \<theta>')(auto)
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   238
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   239
lemma w2:
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   240
  assumes "name |\<in>| dom \<theta>'" 
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   241
  shows "\<theta><lookup \<theta>' name> = lookup (\<theta><\<theta>'>) name"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   242
using assms
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   243
apply(induct \<theta>')
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   244
apply(auto simp add: notin_empty_fset)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   245
done
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   246
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   247
lemma w3:
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   248
  assumes "name |\<in>| dom \<theta>" 
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   249
  shows "lookup \<theta> name = lookup (\<theta>|p) (p \<bullet> name)"
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   250
using assms
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   251
apply(induct \<theta>)
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   252
apply(auto simp add: notin_empty_fset)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   253
done
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   254
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   255
lemma fresh_lookup:
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   256
  fixes X Y::"name"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   257
  and   \<theta>::"Subst"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   258
  assumes asms: "atom X \<sharp> Y" "atom X \<sharp> \<theta>"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   259
  shows "atom X \<sharp> (lookup \<theta> Y)"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   260
  using assms
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   261
  apply (induct \<theta>)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   262
  apply (auto simp add: fresh_Cons fresh_Pair fresh_at_base ty.fresh)
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   263
  done
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   264
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   265
lemma test:
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   266
  fixes \<theta> \<theta>'::"Subst"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   267
  and T::"ty"
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   268
  assumes "(p \<bullet> atom ` fset (dom \<theta>')) \<sharp>* \<theta>" "supp All [dom \<theta>'].T \<sharp>* p"
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   269
  shows "\<theta><\<theta>'<T>> = \<theta><\<theta>'|p><\<theta><p \<bullet> T>>"
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   270
using assms
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   271
apply(induct T rule: ty.induct)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   272
defer
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   273
apply(auto simp add: tys.supp ty.supp fresh_star_def)[1]
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   274
apply(auto simp add: tys.supp ty.supp fresh_star_def supp_at_base)[1]
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   275
apply(case_tac "name |\<in>| dom \<theta>'")
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   276
apply(subgoal_tac "atom (p \<bullet> name) \<sharp> \<theta>")
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   277
apply(subst (2) lookup_fresh)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   278
apply(perm_simp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   279
apply(auto simp add: fresh_star_def)[1]
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   280
apply(simp)
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   281
apply(simp add: w1)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   282
apply(simp add: w2)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   283
apply(subst w3[symmetric])
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   284
apply(simp add: dom_subst)
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   285
apply(simp)
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   286
apply(perm_simp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   287
apply(rotate_tac 2)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   288
apply(drule_tac p="p" in permute_boolI)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   289
apply(perm_simp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   290
apply(auto simp add: fresh_star_def)[1]
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   291
apply(metis notin_fset)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   292
apply(simp add: lookup_dom)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   293
apply(perm_simp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   294
apply(subst dom_fresh_ty)
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   295
apply(auto)[1]
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   296
apply(rule fresh_lookup)
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   297
apply(simp add: dom_subst)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   298
apply(simp add: dom_pi)
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   299
apply(perm_simp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   300
apply(rotate_tac 2)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   301
apply(drule_tac p="p" in permute_boolI)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   302
apply(perm_simp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   303
apply(simp add: fresh_at_base)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   304
apply (metis in_fset)
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   305
apply(simp add: dom_subst)
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   306
apply(simp add: dom_pi[symmetric])
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   307
apply(perm_simp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   308
apply(subst supp_perm_eq)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   309
apply(simp add: supp_at_base fresh_star_def)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   310
apply (smt Diff_iff atom_eq_iff image_iff insertI1 notin_fset)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   311
apply(simp)
3102
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   312
done
5b5ade6bc889 added definition for generalisation of type schemes (for paper)
Christian Urban <urbanc@in.tum.de>
parents: 3100
diff changeset
   313
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   314
lemma
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   315
  shows "T \<prec>\<prec> S \<Longrightarrow> \<theta><T> \<prec>\<prec> \<theta><S>"
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   316
unfolding generalises_def
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   317
apply(erule exE)+
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   318
apply(erule conjE)+
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   319
apply(rule_tac t="S" and s="All [Xs].T'" in subst)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   320
apply(simp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   321
using at_set_avoiding3
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   322
apply -
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   323
apply(drule_tac x="fset (map_fset atom Xs)" in meta_spec)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   324
apply(drule_tac x="\<theta>" in meta_spec)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   325
apply(drule_tac x="All [Xs].T'" in meta_spec)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   326
apply(drule meta_mp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   327
apply(simp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   328
apply(drule meta_mp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   329
apply(simp add: finite_supp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   330
apply(drule meta_mp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   331
apply(simp add: finite_supp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   332
apply(drule meta_mp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   333
apply(simp add: tys.fresh fresh_star_def)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   334
apply(erule exE)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   335
apply(erule conjE)+
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   336
apply(rule_tac t="All[Xs].T'" and s="p \<bullet> (All [Xs].T')" in subst)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   337
apply(rule supp_perm_eq)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   338
apply(assumption)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   339
apply(perm_simp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   340
apply(subst substs.simps)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   341
apply(perm_simp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   342
apply(simp)
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   343
apply(rule_tac x="\<theta><\<theta>'|p>" in exI)
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   344
apply(rule_tac x="p \<bullet> Xs" in exI)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   345
apply(rule_tac x="\<theta><p \<bullet> T'>" in exI)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   346
apply(rule conjI)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   347
apply(simp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   348
apply(rule conjI)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   349
defer
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   350
apply(simp add: dom_subst)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   351
apply(simp add: dom_pi dom_eqvt[symmetric])
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   352
apply(rule_tac t="T" and s="\<theta>'<T'>" in subst)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   353
apply(simp)
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   354
apply(simp)
3103
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   355
apply(rule test)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   356
apply(perm_simp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   357
apply(rotate_tac 2)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   358
apply(drule_tac p="p" in permute_boolI)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   359
apply(perm_simp)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   360
apply(auto simp add: fresh_star_def)
9a63d90d1752 proved that generalisation is closed under substitution
Christian Urban <urbanc@in.tum.de>
parents: 3102
diff changeset
   361
done
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
3104
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   363
lemma 
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   364
  "T \<prec>\<prec>  All [Xs].T' \<longleftrightarrow> (\<exists>\<theta>. T = \<theta><T'> \<and> dom \<theta> = Xs)"
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   365
apply(auto)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   366
defer
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   367
unfolding generalises_def
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   368
apply(auto)[1]
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   369
apply(auto)[1]
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   370
apply(drule sym)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   371
apply(simp add: Abs_eq_iff2)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   372
apply(simp add: alphas)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   373
apply(auto)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   374
apply(rule_tac x="map (\<lambda>(X, T). (p \<bullet> X, T)) \<theta>" in exI)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   375
apply(auto)
f7c4b8e6918b updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents: 3103
diff changeset
   376
oops
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
end