Nominal/Ex/Foo1.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 19 Aug 2011 11:01:52 +0900
changeset 2995 6d2859aeebba
parent 2950 0911cb7bf696
permissions -rw-r--r--
Comment out examples with 'True' that do not work because function still does not work
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Foo1
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports "../Nominal2" 
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
2564
5be8e34c2c0e tuned example
Christian Urban <urbanc@in.tum.de>
parents: 2562
diff changeset
     5
text {* 
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
  Contrived example that has more than one
2564
5be8e34c2c0e tuned example
Christian Urban <urbanc@in.tum.de>
parents: 2562
diff changeset
     7
  binding function
5be8e34c2c0e tuned example
Christian Urban <urbanc@in.tum.de>
parents: 2562
diff changeset
     8
*}
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
2571
f0252365936c proved that bn functions return a finite set
Christian Urban <urbanc@in.tum.de>
parents: 2564
diff changeset
    10
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
atom_decl name
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
nominal_datatype foo: trm =
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
  Var "name"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
| App "trm" "trm"
2950
0911cb7bf696 changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
parents: 2630
diff changeset
    16
| Lam x::"name" t::"trm"    binds x in t
0911cb7bf696 changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
parents: 2630
diff changeset
    17
| Let1 a::"assg" t::"trm"   binds "bn1 a" in t
0911cb7bf696 changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
parents: 2630
diff changeset
    18
| Let2 a::"assg" t::"trm"   binds "bn2 a" in t
0911cb7bf696 changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
parents: 2630
diff changeset
    19
| Let3 a::"assg" t::"trm"   binds "bn3 a" in t
0911cb7bf696 changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
parents: 2630
diff changeset
    20
| Let4 a::"assg'" t::"trm"  binds (set) "bn4 a" in t
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
and assg =
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  As "name" "name" "trm"
2564
5be8e34c2c0e tuned example
Christian Urban <urbanc@in.tum.de>
parents: 2562
diff changeset
    23
and assg' =
5be8e34c2c0e tuned example
Christian Urban <urbanc@in.tum.de>
parents: 2562
diff changeset
    24
  BNil
5be8e34c2c0e tuned example
Christian Urban <urbanc@in.tum.de>
parents: 2562
diff changeset
    25
| BAs "name" "assg'"
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
binder
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
  bn1::"assg \<Rightarrow> atom list" and
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
  bn2::"assg \<Rightarrow> atom list" and
2564
5be8e34c2c0e tuned example
Christian Urban <urbanc@in.tum.de>
parents: 2562
diff changeset
    29
  bn3::"assg \<Rightarrow> atom list" and
5be8e34c2c0e tuned example
Christian Urban <urbanc@in.tum.de>
parents: 2562
diff changeset
    30
  bn4::"assg' \<Rightarrow> atom set"
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
where
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
  "bn1 (As x y t) = [atom x]"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
| "bn2 (As x y t) = [atom y]"
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
| "bn3 (As x y t) = [atom x, atom y]"
2564
5be8e34c2c0e tuned example
Christian Urban <urbanc@in.tum.de>
parents: 2562
diff changeset
    35
| "bn4 (BNil) = {}"
5be8e34c2c0e tuned example
Christian Urban <urbanc@in.tum.de>
parents: 2562
diff changeset
    36
| "bn4 (BAs a as) = {atom a} \<union> bn4 as"
5be8e34c2c0e tuned example
Christian Urban <urbanc@in.tum.de>
parents: 2562
diff changeset
    37
2598
b136721eedb2 automated permute_bn theorems
Christian Urban <urbanc@in.tum.de>
parents: 2594
diff changeset
    38
thm foo.permute_bn
2593
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
    39
thm foo.perm_bn_alpha
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
    40
thm foo.perm_bn_simps
25dcb2b1329e ordered raw_bn_info to agree with the order of the raw_bn_functions; started alpha_bn proof
Christian Urban <urbanc@in.tum.de>
parents: 2588
diff changeset
    41
thm foo.bn_finite
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
thm foo.distinct
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
thm foo.induct
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
thm foo.inducts
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
thm foo.exhaust
2630
8268b277d240 automated all strong induction lemmas
Christian Urban <urbanc@in.tum.de>
parents: 2627
diff changeset
    47
thm foo.strong_induct
2626
d1bdc281be2b moved all strong_exhaust code to nominal_dt_quot; tuned examples
Christian Urban <urbanc@in.tum.de>
parents: 2598
diff changeset
    48
thm foo.strong_exhaust
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
thm foo.fv_defs
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
thm foo.bn_defs
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
thm foo.perm_simps
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
thm foo.eq_iff
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
thm foo.fv_bn_eqvt
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
thm foo.size_eqvt
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
thm foo.supports
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
thm foo.fsupp
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
thm foo.supp
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
thm foo.fresh
2571
f0252365936c proved that bn functions return a finite set
Christian Urban <urbanc@in.tum.de>
parents: 2564
diff changeset
    59
thm foo.bn_finite
2494
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
end
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
11133eb76f61 added Foo1 to explore a contrived example
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65