src/HOL/Datatype_Examples/Lambda_Term.thy
author wenzelm
Fri, 18 Aug 2017 20:47:47 +0200
changeset 66453 cc19f7ca2ed6
parent 63167 0909deb8059b
permissions -rw-r--r--
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58309
a09ec6daaa19 renamed 'BNF_Examples' to 'Datatype_Examples' (cf. 'datatypes.pdf')
blanchet
parents: 55932
diff changeset
     1
(*  Title:      HOL/Datatype_Examples/Lambda_Term.thy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     2
    Author:     Dmitriy Traytel, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     3
    Author:     Andrei Popescu, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     4
    Copyright   2012
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     5
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     6
Lambda-terms.
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     7
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     8
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
     9
section \<open>Lambda-Terms\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    10
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    11
theory Lambda_Term
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 63167
diff changeset
    12
imports "HOL-Library.FSet"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    13
begin
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    14
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
    15
section \<open>Datatype definition\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    16
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
    17
datatype 'a trm =
49601
ba31032887db modernized examples;
blanchet
parents: 49594
diff changeset
    18
  Var 'a |
ba31032887db modernized examples;
blanchet
parents: 49594
diff changeset
    19
  App "'a trm" "'a trm" |
ba31032887db modernized examples;
blanchet
parents: 49594
diff changeset
    20
  Lam 'a "'a trm" |
ba31032887db modernized examples;
blanchet
parents: 49594
diff changeset
    21
  Lt "('a \<times> 'a trm) fset" "'a trm"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    22
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    23
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 60585
diff changeset
    24
subsection \<open>Example: The set of all variables varsOf and free variables fvarsOf of a term\<close>
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    25
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55129
diff changeset
    26
primrec varsOf :: "'a trm \<Rightarrow> 'a set" where
53355
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    27
  "varsOf (Var a) = {a}"
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    28
| "varsOf (App f x) = varsOf f \<union> varsOf x"
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    29
| "varsOf (Lam x b) = {x} \<union> varsOf b"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 58889
diff changeset
    30
| "varsOf (Lt F t) = varsOf t \<union> (\<Union>{{x} \<union> X | x X. (x,X) |\<in>| fimage (map_prod id varsOf) F})"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    31
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55129
diff changeset
    32
primrec fvarsOf :: "'a trm \<Rightarrow> 'a set" where
53355
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    33
  "fvarsOf (Var x) = {x}"
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    34
| "fvarsOf (App t1 t2) = fvarsOf t1 \<union> fvarsOf t2"
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    35
| "fvarsOf (Lam x t) = fvarsOf t - {x}"
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55530
diff changeset
    36
| "fvarsOf (Lt xts t) = fvarsOf t - {x | x X. (x,X) |\<in>| fimage (map_prod id varsOf) xts} \<union>
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 58889
diff changeset
    37
    (\<Union>{X | x X. (x,X) |\<in>| fimage (map_prod id varsOf) xts})"
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    38
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    39
lemma diff_Un_incl_triv: "\<lbrakk>A \<subseteq> D; C \<subseteq> E\<rbrakk> \<Longrightarrow> A - B \<union> C \<subseteq> D \<union> E" by blast
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55530
diff changeset
    41
lemma in_fimage_map_prod_fset_iff[simp]:
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55530
diff changeset
    42
  "(x, y) |\<in>| fimage (map_prod f g) xts \<longleftrightarrow> (\<exists> t1 t2. (t1, t2) |\<in>| xts \<and> x = f t1 \<and> y = g t2)"
54014
21dac9a60f0c base the fset bnf on the new FSet theory
traytel
parents: 53355
diff changeset
    43
  by force
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    44
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    45
lemma fvarsOf_varsOf: "fvarsOf t \<subseteq> varsOf t"
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
proof induct
53355
603e6e97c391 modernized more examples
traytel
parents: 51804
diff changeset
    47
  case (Lt xts t) thus ?case unfolding fvarsOf.simps varsOf.simps by (elim diff_Un_incl_triv) auto
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    48
qed auto
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    49
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    50
end