Nominal/TySch.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Thu, 18 Mar 2010 23:19:55 +0100
changeset 1525 bf321f16d025
parent 1515 76fa21f27f22
child 1530 24dbe785f2e5
permissions -rw-r--r--
Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
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
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
    13
lemma permute_rsp_fset[quot_respect]:
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
    14
  "(op = ===> op \<approx> ===> op \<approx>) permute permute"
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
    15
  apply (simp add: eqvts[symmetric])
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
    16
  apply clarify
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
  apply (subst permute_minus_cancel(1)[symmetric, of "xb"])
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
    18
  apply (subst mem_eqvt[symmetric])
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
    19
  apply (subst (2) permute_minus_cancel(1)[symmetric, of "xb"])
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
    20
  apply (subst mem_eqvt[symmetric])
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
    21
  apply (erule_tac x="- x \<bullet> xb" in allE)
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
    22
  apply simp
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
  done
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
    24
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
    25
instantiation FSet.fset :: (pt) pt
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
    26
begin
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
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
term "permute :: perm \<Rightarrow> 'a list \<Rightarrow> 'a list"
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
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
quotient_definition
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
  "permute_fset :: perm \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
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
is
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
    33
  "permute :: perm \<Rightarrow> 'a list \<Rightarrow> 'a list"
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
    34
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
    35
lemma permute_list_zero: "0 \<bullet> (x :: 'a list) = x"
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
    36
by (rule permute_zero)
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
    37
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
    38
lemma permute_fset_zero: "0 \<bullet> (x :: 'a fset) = x"
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
    39
by (lifting permute_list_zero)
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
    40
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
    41
lemma permute_list_plus: "(p + q) \<bullet> (x :: 'a list) = p \<bullet> q \<bullet> x"
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
    42
by (rule permute_plus)
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
    43
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
    44
lemma permute_fset_plus: "(p + q) \<bullet> (x :: 'a fset) = p \<bullet> q \<bullet> x"
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
    45
by (lifting permute_list_plus)
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
    46
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
    47
instance
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
    48
apply default
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
    49
apply (rule permute_fset_zero)
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
    50
apply (rule permute_fset_plus)
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
    51
done
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
    52
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
end
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
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
    55
lemma fset_to_set_eqvt[eqvt]: "pi \<bullet> (fset_to_set x) = fset_to_set (pi \<bullet> x)"
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
    56
by (lifting set_eqvt)
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
    57
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
lemma map_eqvt[eqvt]: "pi \<bullet> (map f l) = map (pi \<bullet> f) (pi \<bullet> l)"
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
apply (induct l)
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
apply (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
    61
apply (simp only: eqvt_apply)
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
    62
done
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
    63
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
    64
lemma fmap_eqvt[eqvt]: "pi \<bullet> (fmap f l) = fmap (pi \<bullet> f) (pi \<bullet> l)"
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
    65
by (lifting map_eqvt)
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
    66
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
    67
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
    68
  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
    69
| 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
    70
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
    71
  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
    72
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    73
thm t_tyS.fv
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    74
thm t_tyS.eq_iff
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    75
thm t_tyS.bn
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    76
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
    77
thm t_tyS.inducts
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    78
thm t_tyS.distinct
1430
ccbcebef56f3 Trying to prove atom_image_fresh_swap
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
    79
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
    80
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
    81
  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
    82
  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
    83
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
    84
lemma infinite_Un:
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
    85
  shows "infinite (S \<union> T) \<longleftrightarrow> infinite S \<or> infinite T"
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
    86
  by simp
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
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
    88
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
    89
  shows "fv_t t = supp t" "fv_tyS ts = supp 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
    90
apply(induct rule: t_tyS.inducts)
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
    91
apply(simp_all only: t_tyS.fv)
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
    92
prefer 3
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
    93
apply(rule_tac t="supp (All fset t)" and s="supp (Abs (fset_to_set (fmap atom fset)) t)" in subst)
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
    94
prefer 2
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
apply(subst finite_supp_Abs)
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
    96
apply(drule sym)
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
    97
apply(simp add: finite_fv_t_tyS(1))
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
    98
apply(simp)
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
    99
apply(simp_all (no_asm) only: supp_def)
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
   100
apply(simp_all only: t_tyS.perm)
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
   101
apply(simp_all only: permute_ABS)
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
   102
apply(simp_all only: t_tyS.eq_iff Abs_eq_iff)
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
apply(simp_all only: alpha_gen)
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
   104
apply(simp_all only: eqvts[symmetric])
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
   105
apply(simp_all only: eqvts eqvts_raw)
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
   106
apply(simp_all only: supp_at_base[symmetric,simplified supp_def])
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
   107
apply(simp_all only: infinite_Un[symmetric] Collect_disj_eq[symmetric])
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
   108
apply(simp_all only: de_Morgan_conj[symmetric])
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
   109
done
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
   110
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
   111
instance t and tyS :: fs
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
   112
apply default
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
   113
apply (simp_all add: supp_fv_t_tyS[symmetric] 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
   114
done
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
   115
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
   116
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
   117
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
   118
lemma induct:
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
   119
"\<lbrakk>\<And>name b. P b (Var name);
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
   120
  \<And>t1 t2 b. \<lbrakk>\<And>c. P c t1; \<And>c. P c t2\<rbrakk> \<Longrightarrow> P b (Fun t1 t2);
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
   121
  \<And>fset t. \<lbrakk>\<And>c. P c t; fset_to_set (fmap atom fset) \<sharp>* b\<rbrakk> \<Longrightarrow> P' b (All fset t)
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
   122
 \<rbrakk> \<Longrightarrow> P a t"
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
   123
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
   124
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
   125
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   126
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
   127
  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
   128
  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
   129
  apply(rule_tac x="0::perm" in exI)
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   130
  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
   131
  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
   132
  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
   133
  done
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   134
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   135
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
   136
  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
   137
  apply(simp add: t_tyS.eq_iff)
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   138
  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
   139
  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
   140
  apply auto
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   141
  done
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   142
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   143
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
   144
  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
   145
  apply(simp add: t_tyS.eq_iff)
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   146
  apply(rule_tac x="0::perm" in exI)
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
   147
  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
   148
oops
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   150
lemma
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   151
  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
   152
  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
   153
  using a
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
   154
  apply(simp add: t_tyS.eq_iff)
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   155
  apply(clarify)
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
   156
  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
   157
  apply auto
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   158
  done
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   159
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   160
end