Nominal/Ex/TypeSchemes.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 26 Nov 2011 09:48:14 +0000
changeset 3053 324b148fc6b5
parent 2982 4a00077c008f
child 3071 11f6a561eb4b
permissions -rw-r--r--
used simproc-antiquotation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory TypeSchemes
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
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
section {*** Type Schemes ***}
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
2486
b4ea19604b0b cleaned up two examples
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
     9
(* defined as a single nominal datatype *)
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
nominal_datatype ty =
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
  Var "name"
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
| Fun "ty" "ty"
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
and tys =
2950
0911cb7bf696 changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
parents: 2886
diff changeset
    15
  All xs::"name fset" ty::"ty" binds (set+) xs in ty
2434
92dc6cfa3a95 automatic lifting
Christian Urban <urbanc@in.tum.de>
parents: 2424
diff changeset
    16
2885
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    17
ML {*
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    18
get_all_info @{context}
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    19
*}
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    20
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    21
ML {*
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    22
get_info @{context} "TypeSchemes.ty"
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    23
*}
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    24
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    25
ML {*
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    26
#strong_exhaust (the_info @{context} "TypeSchemes.tys")
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    27
*}
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    28
2468
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    29
thm ty_tys.distinct
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    30
thm ty_tys.induct
2617
e44551d067e6 properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
    31
thm ty_tys.inducts
2885
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    32
thm ty_tys.exhaust 
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
    33
thm ty_tys.strong_exhaust
2468
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    34
thm ty_tys.fv_defs
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    35
thm ty_tys.bn_defs
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    36
thm ty_tys.perm_simps
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    37
thm ty_tys.eq_iff
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    38
thm ty_tys.fv_bn_eqvt
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    39
thm ty_tys.size_eqvt
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
    40
thm ty_tys.supports
2493
2e174807c891 added postprocessed fresh-lemmas for constructors
Christian Urban <urbanc@in.tum.de>
parents: 2486
diff changeset
    41
thm ty_tys.supp
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
    42
thm ty_tys.fresh
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    44
fun
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    45
  lookup :: "(name \<times> ty) list \<Rightarrow> name \<Rightarrow> ty"
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    46
where
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    47
  "lookup [] Y = Var Y"
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    48
| "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
    49
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    50
lemma lookup_eqvt[eqvt]:
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    51
  shows "(p \<bullet> lookup Ts T) = lookup (p \<bullet> Ts) (p \<bullet> T)"
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    52
apply(induct Ts T rule: lookup.induct)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    53
apply(simp_all)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    54
done
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
    55
2981
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
    56
lemma TEST1:
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
    57
  assumes "x = Inl y"
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
    58
  shows "(p \<bullet> Sum_Type.Projl x) = Sum_Type.Projl (p \<bullet> x)"
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
    59
using assms by simp
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
    60
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
    61
lemma TEST2:
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
    62
  assumes "x = Inr y"
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
    63
  shows "(p \<bullet> Sum_Type.Projr x) = Sum_Type.Projr (p \<bullet> x)"
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
    64
using assms by simp
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
    65
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
    66
lemma test:
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
    67
  assumes a: "\<exists>y. f x = Inl y"
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
    68
  shows "(p \<bullet> (Sum_Type.Projl (f x))) = Sum_Type.Projl ((p \<bullet> f) (p \<bullet> x))"
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
    69
using a
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
    70
apply clarify
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
    71
apply(frule_tac p="p" in permute_boolI)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
    72
apply(simp (no_asm_use) only: eqvts)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
    73
apply(subst (asm) permute_fun_app_eq)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
    74
back
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
    75
apply(simp)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
    76
done
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
    77
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
    78
lemma test2:
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
    79
  assumes a: "\<exists>y. f x = Inl y"
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
    80
  shows "(p \<bullet> (Sum_Type.Projl (f x))) = Sum_Type.Projl (p \<bullet> (f x))"
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
    81
using a
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
    82
apply clarify
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
    83
apply(frule_tac p="p" in permute_boolI)
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
    84
apply(simp (no_asm_use) only: eqvts)
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
    85
apply(subst (asm) permute_fun_app_eq)
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
    86
back
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
    87
apply(simp)
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
    88
done
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
    89
2848
da7e6655cd4c fixed the problem when giving a complex default-term; the fundef lemmas in Nominal_Base were not general enough
Christian Urban <urbanc@in.tum.de>
parents: 2843
diff changeset
    90
nominal_primrec (default "sum_case (\<lambda>x. Inl undefined) (\<lambda>x. Inr undefined)")
2838
36544bac1659 More experiments with 'default'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2836
diff changeset
    91
    subst  :: "(name \<times> ty) list \<Rightarrow> ty \<Rightarrow> ty"
36544bac1659 More experiments with 'default'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2836
diff changeset
    92
and substs :: "(name \<times> ty) list \<Rightarrow> tys \<Rightarrow> tys"
36544bac1659 More experiments with 'default'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2836
diff changeset
    93
where
36544bac1659 More experiments with 'default'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2836
diff changeset
    94
  "subst \<theta> (Var X) = lookup \<theta> X"
36544bac1659 More experiments with 'default'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2836
diff changeset
    95
| "subst \<theta> (Fun T1 T2) = Fun (subst \<theta> T1) (subst \<theta> T2)"
36544bac1659 More experiments with 'default'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2836
diff changeset
    96
| "fset (map_fset atom xs) \<sharp>* \<theta> \<Longrightarrow> substs \<theta> (All xs T) = All xs (subst \<theta> T)"
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
    97
(*unfolding subst_substs_graph_def eqvt_def
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
    98
apply rule
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
    99
apply perm_simp
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   100
apply (subst test3)
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   101
defer
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   102
apply (subst test3)
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   103
defer
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   104
apply (subst test3)
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   105
defer
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   106
apply rule*)
2885
1264f2a21ea9 some rudimentary infrastructure for storing data about nominal datatypes
Christian Urban <urbanc@in.tum.de>
parents: 2867
diff changeset
   107
thm subst_substs_graph.intros
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   108
apply(subgoal_tac "\<And>p x r. subst_substs_graph x r \<Longrightarrow> subst_substs_graph (p \<bullet> x) (p \<bullet> r)")
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   109
apply(simp add: eqvt_def)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   110
apply(rule allI)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   111
apply(simp add: permute_fun_def permute_bool_def)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   112
apply(rule ext)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   113
apply(rule ext)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   114
apply(rule iffI)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   115
apply(drule_tac x="p" in meta_spec)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   116
apply(drule_tac x="- p \<bullet> x" in meta_spec)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   117
apply(drule_tac x="- p \<bullet> xa" in meta_spec)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   118
apply(simp)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   119
apply(drule_tac x="-p" in meta_spec)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   120
apply(drule_tac x="x" in meta_spec)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   121
apply(drule_tac x="xa" in meta_spec)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   122
apply(simp)
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   123
--"Eqvt One way"
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   124
apply(erule subst_substs_graph.induct)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   125
apply(perm_simp)
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   126
apply(rule subst_substs_graph.intros)
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   127
apply(erule subst_substs_graph.cases)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   128
apply(simp (no_asm_use) only: eqvts)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   129
apply(subst test)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   130
back
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   131
apply (rule exI)
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   132
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   133
apply(rotate_tac 1)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   134
apply(erule subst_substs_graph.cases)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   135
apply(subst test)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   136
back
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   137
apply (rule exI)
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   138
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   139
apply(perm_simp)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   140
apply(rule subst_substs_graph.intros)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   141
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   142
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   143
apply(subst test)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   144
back
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   145
apply (rule exI)
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   146
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   147
apply(perm_simp)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   148
apply(rule subst_substs_graph.intros)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   149
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   150
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   151
apply(simp)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   152
--"A"
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   153
apply(simp (no_asm_use) only: eqvts)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   154
apply(subst test)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   155
back
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   156
apply (rule exI)
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   157
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   158
apply(rotate_tac 1)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   159
apply(erule subst_substs_graph.cases)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   160
apply(subst test)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   161
back
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   162
apply (rule exI)
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   163
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   164
apply(perm_simp)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   165
apply(rule subst_substs_graph.intros)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   166
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   167
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   168
apply(subst test)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   169
back
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   170
apply (rule exI)
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   171
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   172
apply(perm_simp)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   173
apply(rule subst_substs_graph.intros)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   174
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   175
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   176
apply(simp)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   177
--"A"
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   178
apply(simp)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   179
apply(erule subst_substs_graph.cases)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   180
apply(simp (no_asm_use) only: eqvts)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   181
apply(subst test)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   182
back
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   183
back
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   184
apply (rule exI)
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   185
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   186
apply(rule subst_substs_graph.intros)
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   187
apply (simp add: eqvts)
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   188
apply (subgoal_tac "(p \<bullet> (atom ` fset xs)) \<sharp>* (p \<bullet> \<theta>)")
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   189
apply (simp add: image_eqvt eqvts_raw eqvts)
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   190
apply (simp add: fresh_star_permute_iff)
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   191
apply(perm_simp)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   192
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   193
apply(simp (no_asm_use) only: eqvts)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   194
apply(subst test)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   195
back
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   196
back
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   197
apply (rule exI)
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   198
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   199
apply(rule subst_substs_graph.intros)
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   200
apply (simp add: eqvts)
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   201
apply (subgoal_tac "(p \<bullet> (atom ` fset xs)) \<sharp>* (p \<bullet> \<theta>)")
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   202
apply (simp add: image_eqvt eqvts_raw eqvts)
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   203
apply (simp add: fresh_star_permute_iff)
2709
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   204
apply(perm_simp)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   205
apply(assumption)
eb4a2f4078ae some experiments
Christian Urban <urbanc@in.tum.de>
parents: 2707
diff changeset
   206
apply(simp)
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   207
--"Eqvt done"
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
   208
apply(rule TrueI)
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   209
apply (case_tac x)
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   210
apply simp apply clarify 
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   211
apply (rule_tac y="b" in ty_tys.exhaust(1))
2787
1a6593bc494d added eq_iff and distinct lemmas of nominal datatypes to the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 2728
diff changeset
   212
apply (auto)[1]
1a6593bc494d added eq_iff and distinct lemmas of nominal datatypes to the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 2728
diff changeset
   213
apply (auto)[1]
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   214
apply simp apply clarify 
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   215
apply (rule_tac ya="b" and c="a" in ty_tys.strong_exhaust(2))
2787
1a6593bc494d added eq_iff and distinct lemmas of nominal datatypes to the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 2728
diff changeset
   216
apply (auto)[1]
1a6593bc494d added eq_iff and distinct lemmas of nominal datatypes to the simplifier
Christian Urban <urbanc@in.tum.de>
parents: 2728
diff changeset
   217
apply (auto)[5]
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   218
--"LAST GOAL"
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   219
apply (simp add: meta_eq_to_obj_eq[OF subst_def, symmetric, unfolded fun_eq_iff])
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   220
apply (subgoal_tac "eqvt_at (\<lambda>(l, r). subst l r) (\<theta>', T)")
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   221
apply (thin_tac "eqvt_at subst_substs_sumC (Inl (\<theta>', T))")
2867
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2859
diff changeset
   222
apply (thin_tac "eqvt_at subst_substs_sumC (Inl (\<theta>', Ta))")
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
   223
prefer 2
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   224
apply (simp add: eqvt_at_def subst_def)
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   225
apply rule
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   226
apply (subst test2)
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   227
apply (simp add: subst_substs_sumC_def)
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   228
apply (simp add: THE_default_def)
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   229
apply (case_tac "Ex1 (subst_substs_graph (Inl (\<theta>', T)))")
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   230
prefer 2
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   231
apply simp
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   232
apply (simp add: the1_equality)
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   233
apply auto[1]
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   234
apply (erule_tac x="x" in allE)
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   235
apply simp
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   236
apply(cases rule: subst_substs_graph.cases)
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   237
apply assumption
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   238
apply (rule_tac x="lookup \<theta> X" in exI)
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   239
apply clarify
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   240
apply (rule the1_equality)
2867
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2859
diff changeset
   241
apply blast apply assumption
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   242
apply (rule_tac x="(Fun (Sum_Type.Projl (subst_substs_sum (Inl (\<theta>, T1))))
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   243
                  (Sum_Type.Projl (subst_substs_sum (Inl (\<theta>, T2)))))" in exI)
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   244
apply clarify
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   245
apply (rule the1_equality)
2867
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2859
diff changeset
   246
apply blast apply assumption
2710
7eebe0d5d298 Experiments with functions
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2709
diff changeset
   247
apply clarify
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   248
apply simp
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   249
--"-"
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
   250
apply clarify
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
   251
  apply (frule supp_eqvt_at)
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
   252
  apply (simp add: finite_supp)
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
   253
  apply (erule Abs_res_fcb)
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
   254
  apply (simp add: Abs_fresh_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
   255
  apply (simp add: Abs_fresh_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
   256
  apply auto[1]
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
   257
  apply (simp add: fresh_def fresh_star_def)
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
   258
  apply (erule contra_subsetD)
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
   259
  apply (simp add: supp_Pair)
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
   260
  apply blast
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
   261
  apply clarify
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
   262
  apply (simp)
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
   263
  apply (simp add: eqvt_at_def)
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
   264
  apply (subst Abs_eq_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
   265
  apply (rule_tac x="0::perm" in exI)
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
   266
  apply (subgoal_tac "p \<bullet> \<theta>' = \<theta>'")
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
   267
  apply (simp add: alphas fresh_star_zero)
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
   268
  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")
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
   269
  apply blast
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
   270
  apply (subgoal_tac "x \<in> supp(p \<bullet> \<theta>', p \<bullet> T)")
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
   271
  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
   272
  apply auto[1]
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
   273
  apply (subgoal_tac "(atom ` fset (p \<bullet> xs)) \<sharp>* \<theta>'")
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
   274
  apply (simp add: fresh_star_def fresh_def)
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
   275
  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
   276
  apply (simp add: 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
   277
  apply (simp add: fresh_star_def fresh_def)
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
   278
  apply (simp (no_asm) only: supp_eqvt[symmetric] Pair_eqvt[symmetric])
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
   279
  apply (subgoal_tac "p \<bullet> supp (subst \<theta>' T) \<subseteq> p \<bullet> supp (\<theta>', T)")
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
   280
  apply (erule subsetD)
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
   281
  apply (simp add: supp_eqvt)
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
   282
  apply (metis le_eqvt permute_boolI)
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
   283
  apply (rule perm_supp_eq)
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
   284
  apply (simp add: fresh_def fresh_star_def)
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
   285
  apply blast
2850
354a930ef18f TypeSchemes work with 'default'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2848
diff changeset
   286
  done
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   287
2981
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
   288
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
   289
termination (eqvt) by lexicographic_order
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
   290
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2975
diff changeset
   291
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   292
section {* defined as two separate nominal datatypes *}
2486
b4ea19604b0b cleaned up two examples
Christian Urban <urbanc@in.tum.de>
parents: 2480
diff changeset
   293
2308
387fcbd33820 fixed problem with bn_info
Christian Urban <urbanc@in.tum.de>
parents: 2181
diff changeset
   294
nominal_datatype ty2 =
387fcbd33820 fixed problem with bn_info
Christian Urban <urbanc@in.tum.de>
parents: 2181
diff changeset
   295
  Var2 "name"
387fcbd33820 fixed problem with bn_info
Christian Urban <urbanc@in.tum.de>
parents: 2181
diff changeset
   296
| Fun2 "ty2" "ty2"
387fcbd33820 fixed problem with bn_info
Christian Urban <urbanc@in.tum.de>
parents: 2181
diff changeset
   297
387fcbd33820 fixed problem with bn_info
Christian Urban <urbanc@in.tum.de>
parents: 2181
diff changeset
   298
nominal_datatype tys2 =
2950
0911cb7bf696 changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
parents: 2886
diff changeset
   299
  All2 xs::"name fset" ty::"ty2" binds (set+) xs in ty
2337
b151399bd2c3 fixed according to changes in quotient
Christian Urban <urbanc@in.tum.de>
parents: 2308
diff changeset
   300
2468
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
   301
thm tys2.distinct
2630
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2622
diff changeset
   302
thm tys2.induct tys2.strong_induct
2617
e44551d067e6 properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents: 2611
diff changeset
   303
thm tys2.exhaust tys2.strong_exhaust
2468
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
   304
thm tys2.fv_defs
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
   305
thm tys2.bn_defs
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
   306
thm tys2.perm_simps
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
   307
thm tys2.eq_iff
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
   308
thm tys2.fv_bn_eqvt
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
   309
thm tys2.size_eqvt
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
   310
thm tys2.supports
2493
2e174807c891 added postprocessed fresh-lemmas for constructors
Christian Urban <urbanc@in.tum.de>
parents: 2486
diff changeset
   311
thm tys2.supp
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents: 2493
diff changeset
   312
thm tys2.fresh
2468
7b1470b55936 moved a proof to Abs
Christian Urban <urbanc@in.tum.de>
parents: 2454
diff changeset
   313
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   314
fun
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   315
  lookup2 :: "(name \<times> ty2) list \<Rightarrow> name \<Rightarrow> ty2"
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   316
where
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   317
  "lookup2 [] Y = Var2 Y"
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   318
| "lookup2 ((X, T) # Ts) Y = (if X = Y then T else lookup2 Ts Y)"
2556
8ed62410236e added a test about subtyping; disabled two tests, because of problem with function package
Christian Urban <urbanc@in.tum.de>
parents: 2524
diff changeset
   319
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   320
lemma lookup2_eqvt[eqvt]:
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   321
  shows "(p \<bullet> lookup2 Ts T) = lookup2 (p \<bullet> Ts) (p \<bullet> T)"
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   322
  by (induct Ts T rule: lookup2.induct) simp_all
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   323
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2676
diff changeset
   324
nominal_primrec
2859
2eeb75cccb8d all tests work again
Christian Urban <urbanc@in.tum.de>
parents: 2850
diff changeset
   325
  subst2  :: "(name \<times> ty2) list \<Rightarrow> ty2 \<Rightarrow> ty2"
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   326
where
2859
2eeb75cccb8d all tests work again
Christian Urban <urbanc@in.tum.de>
parents: 2850
diff changeset
   327
  "subst2 \<theta> (Var2 X) = lookup2 \<theta> X"
2eeb75cccb8d all tests work again
Christian Urban <urbanc@in.tum.de>
parents: 2850
diff changeset
   328
| "subst2 \<theta> (Fun2 T1 T2) = Fun2 (subst2 \<theta> T1) (subst2 \<theta> T2)"
2eeb75cccb8d all tests work again
Christian Urban <urbanc@in.tum.de>
parents: 2850
diff changeset
   329
  unfolding eqvt_def subst2_graph_def
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   330
  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
   331
  apply(rule TrueI)
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   332
  apply(case_tac x)
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   333
  apply(rule_tac y="b" in ty2.exhaust)
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   334
  apply(blast)
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   335
  apply(blast)
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   336
  apply(simp_all add: ty2.distinct)
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   337
  done
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   338
2973
d1038e67923a added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
parents: 2950
diff changeset
   339
termination (eqvt)
2859
2eeb75cccb8d all tests work again
Christian Urban <urbanc@in.tum.de>
parents: 2850
diff changeset
   340
  by lexicographic_order
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   341
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   342
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   343
lemma supp_fun_app2_eqvt:
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   344
  assumes e: "eqvt f"
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   345
  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
   346
  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
   347
  by blast
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   348
 
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   349
lemma supp_subst:
2859
2eeb75cccb8d all tests work again
Christian Urban <urbanc@in.tum.de>
parents: 2850
diff changeset
   350
  "supp (subst2 \<theta> t) \<subseteq> supp \<theta> \<union> supp t"
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   351
  apply (rule supp_fun_app2_eqvt)
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   352
  unfolding eqvt_def by (simp add: eqvts_raw)
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   353
 
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   354
lemma fresh_star_inter1:
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   355
  "xs \<sharp>* z \<Longrightarrow> (xs \<inter> ys) \<sharp>* z"
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   356
  unfolding fresh_star_def by blast
2830
297cff1d1048 FCB for res binding and simplified proof of subst for type schemes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2822
diff changeset
   357
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   358
nominal_primrec
2859
2eeb75cccb8d all tests work again
Christian Urban <urbanc@in.tum.de>
parents: 2850
diff changeset
   359
  substs2 :: "(name \<times> ty2) list \<Rightarrow> tys2 \<Rightarrow> tys2"
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   360
where
2859
2eeb75cccb8d all tests work again
Christian Urban <urbanc@in.tum.de>
parents: 2850
diff changeset
   361
  "fset (map_fset atom xs) \<sharp>* \<theta> \<Longrightarrow> substs2 \<theta> (All2 xs t) = All2 xs (subst2 \<theta> t)"
2eeb75cccb8d all tests work again
Christian Urban <urbanc@in.tum.de>
parents: 2850
diff changeset
   362
  unfolding eqvt_def substs2_graph_def
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   363
  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
   364
  apply auto[2]
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   365
  apply (rule_tac y="b" and c="a" in tys2.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
   366
  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
   367
  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
   368
  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
   369
  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
   370
  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
   371
  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
   372
  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
   373
  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
   374
  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
   375
  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
   376
  apply(auto simp add: fresh_star_def fresh_def)[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
   377
  apply(subgoal_tac "x \<in> supp t")
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
   378
  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
   379
  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
   380
  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
   381
  apply(blast)
2832
76db0b854bf6 Simpler proof of TypeSchemes/substs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2831
diff changeset
   382
  apply clarify
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
   383
  apply (simp add: subst2.eqvt)
2801
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   384
  apply (subst Abs_eq_iff)
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   385
  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
   386
  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
   387
  apply (simp add: alphas fresh_star_zero)
2859
2eeb75cccb8d all tests work again
Christian Urban <urbanc@in.tum.de>
parents: 2850
diff changeset
   388
  apply (subgoal_tac "\<And>x. x \<in> supp (subst2 \<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
   389
  apply blast
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
   390
  apply (subgoal_tac "x \<in> supp(p \<bullet> \<theta>', p \<bullet> t)")
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
   391
  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
   392
  apply auto[1]
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
   393
  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
   394
  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
   395
  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
   396
  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
   397
  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
   398
  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
   399
  apply (simp add: supp_Pair)
2832
76db0b854bf6 Simpler proof of TypeSchemes/substs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2831
diff changeset
   400
  apply (rule perm_supp_eq)
76db0b854bf6 Simpler proof of TypeSchemes/substs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2831
diff changeset
   401
  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
   402
  apply blast
5ecb857e9de7 proved subst for All constructor in type schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2787
diff changeset
   403
  done
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   404
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   405
text {* Some Tests about Alpha-Equality *}
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
lemma
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
  shows "All {|a, b|} (Fun (Var a) (Var b)) = All {|b, a|} (Fun (Var a) (Var b))"
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   409
  apply(simp add: ty_tys.eq_iff Abs_eq_iff)
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
  apply(rule_tac x="0::perm" in exI)
2676
028d5511c15f some tryes about substitution over type-schemes
Christian Urban <urbanc@in.tum.de>
parents: 2634
diff changeset
   411
  apply(simp add: alphas fresh_star_def ty_tys.supp supp_at_base)
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
  done
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
lemma
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
  shows "All {|a, b|} (Fun (Var a) (Var b)) = All {|a, b|} (Fun (Var b) (Var a))"
2566
a59d8e1e3a17 moved rest of the lemmas from Nominal2_FSet to the TypeScheme example
Christian Urban <urbanc@in.tum.de>
parents: 2556
diff changeset
   416
  apply(simp add: ty_tys.eq_iff Abs_eq_iff)
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
  apply(rule_tac x="(atom a \<rightleftharpoons> atom b)" in exI)
2566
a59d8e1e3a17 moved rest of the lemmas from Nominal2_FSet to the TypeScheme example
Christian Urban <urbanc@in.tum.de>
parents: 2556
diff changeset
   418
  apply(simp add: alphas fresh_star_def supp_at_base ty_tys.supp)
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
  done
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
lemma
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   422
  shows "All {|a, b, c|} (Fun (Var a) (Var b)) = All {|a, b|} (Fun (Var a) (Var b))"
2566
a59d8e1e3a17 moved rest of the lemmas from Nominal2_FSet to the TypeScheme example
Christian Urban <urbanc@in.tum.de>
parents: 2556
diff changeset
   423
  apply(simp add: ty_tys.eq_iff Abs_eq_iff)
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
  apply(rule_tac x="0::perm" in exI)
2566
a59d8e1e3a17 moved rest of the lemmas from Nominal2_FSet to the TypeScheme example
Christian Urban <urbanc@in.tum.de>
parents: 2556
diff changeset
   425
  apply(simp add: alphas fresh_star_def ty_tys.supp supp_at_base)
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
done
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
lemma
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
  assumes a: "a \<noteq> b"
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
  shows "\<not>(All {|a, b|} (Fun (Var a) (Var b)) = All {|c|} (Fun (Var c) (Var c)))"
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
  using a
2566
a59d8e1e3a17 moved rest of the lemmas from Nominal2_FSet to the TypeScheme example
Christian Urban <urbanc@in.tum.de>
parents: 2556
diff changeset
   432
  apply(simp add: ty_tys.eq_iff Abs_eq_iff)
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
  apply(clarify)
2566
a59d8e1e3a17 moved rest of the lemmas from Nominal2_FSet to the TypeScheme example
Christian Urban <urbanc@in.tum.de>
parents: 2556
diff changeset
   434
  apply(simp add: alphas fresh_star_def ty_tys.eq_iff ty_tys.supp supp_at_base)
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
  apply auto
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
  done
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
2566
a59d8e1e3a17 moved rest of the lemmas from Nominal2_FSet to the TypeScheme example
Christian Urban <urbanc@in.tum.de>
parents: 2556
diff changeset
   438
1795
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
e39453c8b186 tuned type-schemes example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
end