Nominal/TySch.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 19 Mar 2010 18:42:57 +0100
changeset 1553 4355eb3b7161
parent 1547 57f7af5d7564
child 1561 c3dca6e600c8
permissions -rw-r--r--
Automatically derive support for datatypes with at-most one binding per constructor.
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 *}
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
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
    11
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
    12
  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
    13
| 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
    14
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
    15
  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
    16
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    17
thm t_tyS.fv
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    18
thm t_tyS.eq_iff
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    19
thm t_tyS.bn
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    20
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
    21
thm t_tyS.inducts
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    22
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
    23
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
    24
1553
4355eb3b7161 Automatically derive support for datatypes with at-most one binding per constructor.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1547
diff changeset
    25
lemmas t_tyS_supp = t_tyS.fv[simplified t_tyS.supp]
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
    26
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 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
    28
  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
    29
  and     a2: "\<And>t1 t2 b. \<lbrakk>\<And>c. P c t1; \<And>c. P c t2\<rbrakk> \<Longrightarrow> P b (Fun t1 t2)"
1538
6853ce305118 Working on TySch strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1537
diff changeset
    30
  and     a3: "\<And>fset t b. \<lbrakk>\<And>c. P c t; fset_to_set (fmap atom fset) \<sharp>* b\<rbrakk> \<Longrightarrow> P' b (All fset t)"
1539
78d0adf8a086 TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1538
diff changeset
    31
  shows "P (a :: 'a :: pt) t \<and> P' d ts "
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
    32
proof -
1539
78d0adf8a086 TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1538
diff changeset
    33
  have " (\<forall>p a. P a (p \<bullet> t)) \<and> (\<forall>p d. P' d (p \<bullet> ts))"
1538
6853ce305118 Working on TySch strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1537
diff changeset
    34
    apply (rule t_tyS.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
    35
    apply (simp add: a1)
1539
78d0adf8a086 TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1538
diff changeset
    36
    apply (simp)
78d0adf8a086 TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1538
diff changeset
    37
    apply (rule allI)+
1538
6853ce305118 Working on TySch strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1537
diff changeset
    38
    apply (rule a2)
1539
78d0adf8a086 TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1538
diff changeset
    39
    apply simp
78d0adf8a086 TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1538
diff changeset
    40
    apply simp
78d0adf8a086 TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1538
diff changeset
    41
    apply (rule allI)
1538
6853ce305118 Working on TySch strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1537
diff changeset
    42
    apply (rule allI)
1539
78d0adf8a086 TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1538
diff changeset
    43
    apply(subgoal_tac "\<exists>new::name fset. fset_to_set (fmap atom new) \<sharp>* (d, All (p \<bullet> fset) (p \<bullet> t))
78d0adf8a086 TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1538
diff changeset
    44
                                      \<and> fcard new = fcard fset")
78d0adf8a086 TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1538
diff changeset
    45
    apply clarify
78d0adf8a086 TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1538
diff changeset
    46
    (*apply(rule_tac t="p \<bullet> All fset t" and 
78d0adf8a086 TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1538
diff changeset
    47
                   s="(((p \<bullet> fset) \<leftrightarrow> new) + p) \<bullet> All fset t" in subst)
1538
6853ce305118 Working on TySch strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1537
diff changeset
    48
    apply (rule a3)
1539
78d0adf8a086 TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1538
diff changeset
    49
    apply simp_all*)
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
    50
    sorry
1538
6853ce305118 Working on TySch strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1537
diff changeset
    51
  then have "P a (0 \<bullet> t) \<and> P' d (0 \<bullet> ts)" by blast
6853ce305118 Working on TySch strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1537
diff changeset
    52
  then show ?thesis by simp
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
    53
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
    54
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    55
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
    56
  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
    57
  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
    58
  apply(rule_tac x="0::perm" in exI)
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    59
  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
    60
  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
    61
  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
    62
  done
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    63
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    64
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
    65
  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
    66
  apply(simp add: t_tyS.eq_iff)
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    67
  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
    68
  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
    69
  apply auto
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
  done
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    71
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    72
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
    73
  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
    74
  apply(simp add: t_tyS.eq_iff)
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    75
  apply(rule_tac x="0::perm" in exI)
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    76
  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
    77
oops
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    78
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    79
lemma
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    80
  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
    81
  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
    82
  using a
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    83
  apply(simp add: t_tyS.eq_iff)
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
  apply(clarify)
1515
76fa21f27f22 Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1510
diff changeset
    85
  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
    86
  apply auto
1271
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    87
  done
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    88
393aced4801d Forgot to add one file.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
end