Nominal/TySch.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 19 Mar 2010 09:03:10 +0100
changeset 1537 0b21101157b1
parent 1534 984ea1299cd7
child 1538 6853ce305118
permissions -rw-r--r--
Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory TySch
1525
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
     2
imports "Parser" "../Attic/Prove" "FSet"
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
atom_decl name
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
1477
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
     7
ML {* val _ = cheat_fv_rsp := false *}
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
     8
ML {* val _ = cheat_const_rsp := false *}
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
     9
ML {* val _ = cheat_equivp := false *}
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
    10
ML {* val _ = cheat_fv_eqvt := false *}
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
    11
ML {* val _ = cheat_alpha_eqvt := false *}
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
1477
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
    13
nominal_datatype t =
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
    14
  Var "name"
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
    15
| Fun "t" "t"
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
    16
and tyS =
1525
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    17
  All xs::"name fset" ty::"t" bind xs in ty
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    19
thm t_tyS.fv
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    20
thm t_tyS.eq_iff
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    21
thm t_tyS.bn
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    22
thm t_tyS.perm
1525
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    23
thm t_tyS.inducts
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    24
thm t_tyS.distinct
1534
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    25
ML {* Sign.of_sort @{theory} (@{typ t}, @{sort fs}) *}
1430
ccbcebef56f3 Trying to prove atom_image_fresh_swap
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    26
1525
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    27
lemma finite_fv_t_tyS:
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    28
  shows "finite (fv_t t)" "finite (fv_tyS ts)"
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    29
  by (induct rule: t_tyS.inducts) (simp_all)
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    30
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    31
lemma supp_fv_t_tyS:
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    32
  shows "fv_t t = supp t" "fv_tyS ts = supp ts"
1534
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    33
  apply(induct rule: t_tyS.inducts)
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    34
  apply(simp_all only: t_tyS.fv)
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    35
  prefer 3
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    36
  apply(rule_tac t="supp (All fset t)" and s="supp (Abs (fset_to_set (fmap atom fset)) t)" in subst)
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    37
  prefer 2
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    38
  apply(subst finite_supp_Abs)
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    39
  apply(drule sym)
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    40
  apply(simp add: finite_fv_t_tyS(1))
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    41
  apply(simp)
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    42
  apply(simp_all (no_asm) only: supp_def)
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    43
  apply(simp_all only: t_tyS.perm)
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    44
  apply(simp_all only: permute_ABS)
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    45
  apply(simp_all only: t_tyS.eq_iff Abs_eq_iff)
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    46
  apply(simp_all only: alpha_gen)
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    47
  apply(simp_all only: eqvts[symmetric])
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    48
  apply(simp_all only: eqvts eqvts_raw)
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    49
  apply(simp_all only: supp_at_base[symmetric,simplified supp_def])
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    50
  apply(simp_all only: infinite_Un[symmetric] Collect_disj_eq[symmetric])
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    51
  apply(simp_all only: de_Morgan_conj[symmetric])
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    52
  done
1525
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    53
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    54
instance t and tyS :: fs
1534
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    55
  apply default
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    56
  apply (simp_all add: supp_fv_t_tyS[symmetric] finite_fv_t_tyS)
984ea1299cd7 The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1530
diff changeset
    57
  done
1525
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    58
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    59
lemmas t_tyS_supp = t_tyS.fv[simplified supp_fv_t_tyS]
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    60
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    61
lemma induct:
1537
0b21101157b1 Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1534
diff changeset
    62
  assumes a1: "\<And>name b. P b (Var name)"
0b21101157b1 Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1534
diff changeset
    63
  and     a2: "\<And>t1 t2 b. \<lbrakk>\<And>c. P c t1; \<And>c. P c t2\<rbrakk> \<Longrightarrow> P b (Fun t1 t2)"
0b21101157b1 Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1534
diff changeset
    64
  and     a3: "\<And>fset t. \<lbrakk>\<And>c. P c t; fset_to_set (fmap atom fset) \<sharp>* b\<rbrakk> \<Longrightarrow> P' b (All fset t)"
0b21101157b1 Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1534
diff changeset
    65
  shows "P a t"
0b21101157b1 Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1534
diff changeset
    66
proof -
0b21101157b1 Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1534
diff changeset
    67
  have "\<And>p. P a (p \<bullet> t)"
0b21101157b1 Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1534
diff changeset
    68
    apply (induct t rule: t_tyS.inducts(1))
0b21101157b1 Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1534
diff changeset
    69
    apply (simp add: a1)
0b21101157b1 Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1534
diff changeset
    70
    apply (simp_all)
0b21101157b1 Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1534
diff changeset
    71
    apply (rule_tac ?t1.0="p \<bullet> t1" and ?t2.0="p \<bullet> t2" in a2)
0b21101157b1 Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1534
diff changeset
    72
    sorry
0b21101157b1 Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1534
diff changeset
    73
  then have "P a (0 \<bullet> t)" by blast
0b21101157b1 Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1534
diff changeset
    74
  then show "P a t" by simp
0b21101157b1 Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1534
diff changeset
    75
qed
1525
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    76
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    77
lemma
1525
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    78
  shows "All {|a, b|} (Fun (Var a) (Var b)) = All {|b, a|} (Fun (Var a) (Var b))"
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    79
  apply(simp add: t_tyS.eq_iff)
1477
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
    80
  apply(rule_tac x="0::perm" in exI)
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    81
  apply(simp add: alpha_gen)
1477
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
    82
  apply(auto)
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
    83
  apply(simp add: fresh_star_def fresh_zero_perm)
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
  done
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    85
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    86
lemma
1525
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    87
  shows "All {|a, b|} (Fun (Var a) (Var b)) = All {|a, b|} (Fun (Var b) (Var a))"
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    88
  apply(simp add: t_tyS.eq_iff)
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
  apply(rule_tac x="(atom a \<rightleftharpoons> atom b)" in exI)
1510
be911e869fde Added fv,bn,distinct,perm to the simplifier.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1498
diff changeset
    90
  apply(simp add: alpha_gen fresh_star_def eqvts)
1477
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
    91
  apply auto
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    92
  done
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    94
lemma
1525
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
    95
  shows "All {|a, b, c|} (Fun (Var a) (Var b)) = All {|a, b|} (Fun (Var a) (Var b))"
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    96
  apply(simp add: t_tyS.eq_iff)
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    97
  apply(rule_tac x="0::perm" in exI)
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    98
  apply(simp add: alpha_gen fresh_star_def eqvts t_tyS.eq_iff)
1477
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
    99
oops
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
lemma
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   102
  assumes a: "a \<noteq> b"
1525
bf321f16d025 Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1515
diff changeset
   103
  shows "\<not>(All {|a, b|} (Fun (Var a) (Var b)) = All {|c|} (Fun (Var c) (Var c)))"
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
  using a
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
   105
  apply(simp add: t_tyS.eq_iff)
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   106
  apply(clarify)
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
   107
  apply(simp add: alpha_gen fresh_star_def eqvts t_tyS.eq_iff)
1477
4ac3485899e1 Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1430
diff changeset
   108
  apply auto
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   109
  done
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   110
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   111
end